2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory LamTest
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Nominal2"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
nominal_datatype lam =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
Var "name"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
| App "lam" "lam"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
| Lam x::"name" l::"lam" bind x in l
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
2651
|
12 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
val trace = Unsynchronized.ref false
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
fun trace_msg msg = if ! trace then tracing (msg ()) else ()
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
val boolT = HOLogic.boolT
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
val mk_eq = HOLogic.mk_eq
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
open Function_Lib
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
open Function_Common
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
datatype globals = Globals of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
{fvar: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
domT: typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
ranT: typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
h: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
y: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
x: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
z: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
a: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
P: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
D: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
Pbool:term}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
datatype rec_call_info = RCInfo of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
{RIvs: (string * typ) list, (* Call context: fixes and assumes *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
CCas: thm list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
rcarg: term, (* The recursive argument *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
llRI: thm,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
h_assum: term}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
datatype clause_context = ClauseContext of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
{ctxt : Proof.context,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
qs : term list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
gs : term list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
lhs: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
rhs: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
cqs: cterm list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
ags: thm list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
case_hyp : thm}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
ClauseContext { ctxt = ProofContext.transfer thy ctxt,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
datatype clause_info = ClauseInfo of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
{no: int,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
qglr : ((string * typ) list * term list * term * term),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
cdata : clause_context,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
tree: Function_Ctx_Tree.ctx_tree,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
lGI: thm,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
RCs: rec_call_info list}
|
2651
|
68 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
|
2651
|
70 |
thm accp_induct_rule
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
|
2651
|
72 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
(* Theory dependencies. *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
val acc_induct_rule = @{thm accp_induct_rule}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
val acc_downward = @{thm accp_downward}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
val accI = @{thm accp.accI}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
val case_split = @{thm HOL.case_split}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
val fundef_default_value = @{thm FunDef.fundef_default_value}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
val not_acc_down = @{thm not_accp_down}
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
85 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
88 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
fun find_calls tree =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
90 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
([], (fixes, assumes, arg) :: xs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
| add_Ri _ _ _ _ = raise Match
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
95 |
rev (Function_Ctx_Tree.traverse_tree add_Ri tree [])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
96 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
99 |
ML {*
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
100 |
fun mk_eqvt_at (f_trm, arg_trm) =
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
101 |
let
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
102 |
val f_ty = fastype_of f_trm
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
103 |
val arg_ty = domain_type f_ty
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
104 |
in
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
105 |
Const (@{const_name eqvt_at}, [f_ty, arg_ty] ---> @{typ bool}) $ f_trm $ arg_trm
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
106 |
|> HOLogic.mk_Trueprop
|
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
107 |
end
|
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
108 |
*}
|
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
109 |
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
110 |
ML {*
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
111 |
fun find_calls2 f t =
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
112 |
let
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
113 |
fun aux (g $ arg) = aux g #> aux arg #> (if f = g then cons ((g, arg)) else I)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
114 |
| aux (Abs (_, _, t)) = aux t
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
115 |
| aux _ = I
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
116 |
in
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
117 |
aux t []
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
118 |
end
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
119 |
*}
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
120 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
123 |
(** building proof obligations *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
125 |
fun mk_compat_proof_obligations domT ranT fvar f glrs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
126 |
let
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
127 |
fun mk_impl ((qs, gs, lhs, rhs), (qs', gs', lhs', rhs')) =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
val shift = incr_boundvars (length qs')
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
130 |
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
131 |
val RCs_rhs = find_calls2 fvar rhs
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
132 |
val RCs_rhs' = find_calls2 fvar rhs'
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
Logic.mk_implies
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
(HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
|> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
138 |
|> fold_rev (curry Logic.mk_implies) (map (shift o mk_eqvt_at) RCs_rhs) (* HERE *)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
139 |
(*|> fold_rev (curry Logic.mk_implies) (map mk_eqvt_at RCs_rhs')*) (* HERE *)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
140 |
|> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
141 |
|> curry abstract_over fvar
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
|> curry subst_bound f
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
143 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
144 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
145 |
map mk_impl (unordered_pairs glrs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
150 |
fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
151 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
153 |
HOLogic.mk_Trueprop Pbool
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
|> fold_rev (curry Logic.mk_implies) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
HOLogic.mk_Trueprop Pbool
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
|> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
|> mk_forall_rename ("x", x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
|> mk_forall_rename ("P", Pbool)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
162 |
end
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
163 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
(** making a context with it's own local bindings **)
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
166 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
171 |
|>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
val thy = ProofContext.theory_of ctxt'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
fun inst t = subst_bounds (rev qs, t)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
val gs = map inst pre_gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
val lhs = inst pre_lhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
val rhs = inst pre_rhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
val cqs = map (cterm_of thy) qs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
val ags = map (Thm.assume o cterm_of thy) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
val case_hyp = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
cqs = cqs, ags = ags, case_hyp = case_hyp }
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
187 |
end
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
188 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
189 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
190 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
191 |
(* lowlevel term function. FIXME: remove *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
fun abstract_over_list vs body =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
194 |
fun abs lev v tm =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
if v aconv tm then Bound lev
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
196 |
else
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
197 |
(case tm of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
199 |
| t $ u => abs lev v t $ abs lev v u
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
200 |
| t => t)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
202 |
fold_index (fn (i, v) => fn t => abs i v t) vs body
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
203 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
204 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
205 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
206 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
207 |
fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
209 |
val Globals {h, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
210 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
211 |
val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
212 |
val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
213 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
214 |
(* Instantiate the GIntro thm with "f" and import into the clause context. *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
215 |
val lGI = GIntro_thm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
216 |
|> Thm.forall_elim (cert f)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
217 |
|> fold Thm.forall_elim cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
218 |
|> fold Thm.elim_implies ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
219 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
220 |
fun mk_call_info (rcfix, rcassm, rcarg) RI =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
221 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
222 |
val llRI = RI
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
223 |
|> fold Thm.forall_elim cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
224 |
|> fold (Thm.forall_elim o cert o Free) rcfix
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
225 |
|> fold Thm.elim_implies ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
226 |
|> fold Thm.elim_implies rcassm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
227 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
228 |
val h_assum =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
229 |
HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
230 |
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
231 |
|> fold_rev (Logic.all o Free) rcfix
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
|> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
|> abstract_over_list (rev qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
236 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
238 |
val RC_infos = map2 mk_call_info RCs RIntro_thms
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
ClauseInfo {no=no, cdata=cdata, qglr=qglr, lGI=lGI, RCs=RC_infos,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
tree=tree}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
243 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
244 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
245 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
246 |
fun store_compat_thms 0 thms = []
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
247 |
| store_compat_thms n thms =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
248 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
249 |
val (thms1, thms2) = chop n thms
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
250 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
251 |
(thms1 :: store_compat_thms (n - 1) thms2)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
252 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
253 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
254 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
255 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
256 |
(* expects i <= j *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
257 |
fun lookup_compat_thm i j cts =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
258 |
nth (nth cts (i - 1)) (j - i)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
(* if j < i, then turn around *)
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
262 |
fun get_compat_thm thy cts eqvtsi eqvtsj i j ctxi ctxj =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
265 |
val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
266 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
267 |
val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
268 |
in if j < i then
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
269 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
val compat = lookup_compat_thm j i cts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
272 |
compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
273 |
|> fold Thm.forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
274 |
|> fold Thm.elim_implies (rev eqvtsj) (* HERE *)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
275 |
|> fold Thm.elim_implies agsj
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
276 |
|> fold Thm.elim_implies agsi
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
277 |
|> Thm.elim_implies ((Thm.assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
278 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
279 |
else
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
281 |
val compat = lookup_compat_thm i j cts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
282 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
283 |
compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
284 |
|> fold Thm.forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
|
2661
|
285 |
|> fold Thm.elim_implies (rev eqvtsi) (* HERE *)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
286 |
|> fold Thm.elim_implies agsi
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
287 |
|> fold Thm.elim_implies agsj
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
288 |
|> Thm.elim_implies (Thm.assume lhsi_eq_lhsj)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
289 |
|> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
290 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
291 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
292 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
293 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
294 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
295 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
296 |
(* Generates the replacement lemma in fully quantified form. *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
297 |
fun mk_replacement_lemma thy h ih_elim clause =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
298 |
let
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
299 |
val ClauseInfo {cdata=ClauseContext {qs, cqs, ags, case_hyp, ...},
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
300 |
RCs, tree, ...} = clause
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
301 |
local open Conv in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
302 |
val ih_conv = arg1_conv o arg_conv o arg_conv
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
303 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
304 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
305 |
val ih_elim_case =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
306 |
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
307 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
308 |
val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
309 |
val h_assums = map (fn RCInfo {h_assum, ...} =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
310 |
Thm.assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
311 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
312 |
val (eql, _) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
313 |
Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
314 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
315 |
val replace_lemma = (eql RS meta_eq_to_obj_eq)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
316 |
|> Thm.implies_intr (cprop_of case_hyp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
317 |
|> fold_rev (Thm.implies_intr o cprop_of) h_assums
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
318 |
|> fold_rev (Thm.implies_intr o cprop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
319 |
|> fold_rev Thm.forall_intr cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
320 |
|> Thm.close_derivation
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
321 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
322 |
replace_lemma
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
323 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
324 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
325 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
326 |
ML {*
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
327 |
fun mk_eqvt_lemma thy ih_eqvt clause =
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
328 |
let
|
2661
|
329 |
val ClauseInfo {cdata=ClauseContext {cqs, ags, case_hyp, ...}, RCs, ...} = clause
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
330 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
331 |
local open Conv in
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
332 |
val ih_conv = arg1_conv o arg_conv o arg_conv
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
333 |
end
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
334 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
335 |
val ih_eqvt_case =
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
336 |
Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_eqvt
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
337 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
338 |
fun prep_eqvt (RCInfo {llRI, RIvs, CCas, ...}) =
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
339 |
(llRI RS ih_eqvt_case)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
340 |
|> fold_rev (Thm.implies_intr o cprop_of) CCas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
341 |
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
342 |
in
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
343 |
map prep_eqvt RCs
|
2661
|
344 |
|> map (fold_rev (Thm.implies_intr o cprop_of) ags)
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
345 |
|> map (Thm.implies_intr (cprop_of case_hyp))
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
346 |
|> map (fold_rev Thm.forall_intr cqs)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
347 |
|> map (Thm.close_derivation)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
348 |
end
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
349 |
*}
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
350 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
351 |
ML {*
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
352 |
fun mk_uniqueness_clause thy globals compat_store eqvts clausei clausej RLj =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
353 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
354 |
val Globals {h, y, x, fvar, ...} = globals
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
355 |
val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, cqs = cqsi,
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
356 |
ags = agsi, ...}, ...} = clausei
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
357 |
val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
358 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
359 |
val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...} =
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
360 |
mk_clause_context x ctxti cdescj
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
361 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
362 |
val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
363 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
364 |
val Ghsj' = map
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
365 |
(fn RCInfo {h_assum, ...} => Thm.assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
366 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
367 |
val y_eq_rhsj'h = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
368 |
val lhsi_eq_lhsj' = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj'))))
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
369 |
(* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
370 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
371 |
val case_hypj' = trans OF [case_hyp, lhsi_eq_lhsj']
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
372 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
373 |
val RLj_import = RLj
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
374 |
|> fold Thm.forall_elim cqsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
375 |
|> fold Thm.elim_implies agsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
376 |
|> fold Thm.elim_implies Ghsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
377 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
378 |
val eqvtsi = nth eqvts (i - 1)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
379 |
|> map (fold Thm.forall_elim cqsi)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
380 |
|> map (fold Thm.elim_implies [case_hyp])
|
2661
|
381 |
|> map (fold Thm.elim_implies agsi)
|
|
382 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
383 |
val eqvtsj = nth eqvts (j - 1)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
384 |
|> map (fold Thm.forall_elim cqsj')
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
385 |
|> map (fold Thm.elim_implies [case_hypj'])
|
2661
|
386 |
|> map (fold Thm.elim_implies agsj')
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
387 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
388 |
val compat = get_compat_thm thy compat_store eqvtsi eqvtsj i j cctxi cctxj
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
389 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
390 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
391 |
(trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
392 |
|> Thm.implies_elim RLj_import
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
393 |
(* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
394 |
|> (fn it => trans OF [it, compat])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
395 |
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
396 |
|> (fn it => trans OF [y_eq_rhsj'h, it])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
397 |
(* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
398 |
|> fold_rev (Thm.implies_intr o cprop_of) Ghsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
399 |
|> fold_rev (Thm.implies_intr o cprop_of) agsj'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
400 |
(* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
401 |
|> Thm.implies_intr (cprop_of y_eq_rhsj'h)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
402 |
|> Thm.implies_intr (cprop_of lhsi_eq_lhsj')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
403 |
|> fold_rev Thm.forall_intr (cterm_of thy h :: cqsj')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
404 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
405 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
406 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
407 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
408 |
ML {*
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
409 |
fun mk_uniqueness_case thy globals G f ihyp ih_intro G_cases compat_store clauses replems eqvtlems clausei =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
410 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
411 |
val Globals {x, y, ranT, fvar, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
412 |
val ClauseInfo {cdata = ClauseContext {lhs, rhs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
413 |
val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
414 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
415 |
val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
416 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
417 |
fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) =
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
418 |
(llRI RS ih_intro_case)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
419 |
|> fold_rev (Thm.implies_intr o cprop_of) CCas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
420 |
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
421 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
422 |
val existence = fold (curry op COMP o prep_RC) RCs lGI
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
423 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
424 |
val P = cterm_of thy (mk_eq (y, rhsC))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
425 |
val G_lhs_y = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
426 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
427 |
val unique_clauses =
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
428 |
map2 (mk_uniqueness_clause thy globals compat_store eqvtlems clausei) clauses replems
|
2661
|
429 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
430 |
fun elim_implies_eta A AB =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
431 |
Thm.compose_no_flatten true (A, 0) 1 AB |> Seq.list_of |> the_single
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
432 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
433 |
val uniqueness = G_cases
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
434 |
|> Thm.forall_elim (cterm_of thy lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
435 |
|> Thm.forall_elim (cterm_of thy y)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
436 |
|> Thm.forall_elim P
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
437 |
|> Thm.elim_implies G_lhs_y
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
438 |
|> fold elim_implies_eta unique_clauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
439 |
|> Thm.implies_intr (cprop_of G_lhs_y)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
440 |
|> Thm.forall_intr (cterm_of thy y)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
441 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
442 |
val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
443 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
444 |
val exactly_one =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
445 |
ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
446 |
|> curry (op COMP) existence
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
447 |
|> curry (op COMP) uniqueness
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
448 |
|> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
449 |
|> Thm.implies_intr (cprop_of case_hyp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
450 |
|> fold_rev (Thm.implies_intr o cprop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
451 |
|> fold_rev Thm.forall_intr cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
452 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
453 |
val function_value =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
454 |
existence
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
455 |
|> Thm.implies_intr ihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
456 |
|> Thm.implies_intr (cprop_of case_hyp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
457 |
|> Thm.forall_intr (cterm_of thy x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
458 |
|> Thm.forall_elim (cterm_of thy lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
459 |
|> curry (op RS) refl
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
460 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
461 |
(exactly_one, function_value)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
462 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
463 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
464 |
|
2655
1c3ad1256f16
instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
465 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
466 |
ML {*
|
2655
1c3ad1256f16
instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
467 |
fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim G_eqvt f_def =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
468 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
469 |
val Globals {h, domT, ranT, x, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
470 |
val thy = ProofContext.theory_of ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
471 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
472 |
(* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
473 |
val ihyp = Term.all domT $ Abs ("z", domT,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
474 |
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
475 |
HOLogic.mk_Trueprop (Const (@{const_name Ex1}, (ranT --> boolT) --> boolT) $
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
476 |
Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
477 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
478 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
479 |
val ihyp_thm = Thm.assume ihyp |> Thm.forall_elim_vars 0
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
480 |
val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
481 |
val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
482 |
|> instantiate' [] [NONE, SOME (cterm_of thy h)]
|
2655
1c3ad1256f16
instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
483 |
val ih_eqvt = ihyp_thm RS (G_eqvt RS (f_def RS @{thm fundef_ex1_eqvt_at}))
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
484 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
485 |
val _ = trace_msg (K "Proving Replacement lemmas...")
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
486 |
val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
487 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
488 |
val _ = trace_msg (K "Proving Equivariance lemmas...")
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
489 |
val eqvtLemmas = map (mk_eqvt_lemma thy ih_eqvt) clauses
|
2651
|
490 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
491 |
val _ = trace_msg (K "Proving cases for unique existence...")
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
492 |
val (ex1s, values) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
493 |
split_list (map (mk_uniqueness_case thy globals G f
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
494 |
ihyp ih_intro G_elim compat_store clauses repLemmas eqvtLemmas) clauses)
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
495 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
496 |
val _ = trace_msg (K "Proving: Graph is a function")
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
497 |
val graph_is_function = complete
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
498 |
|> Thm.forall_elim_vars 0
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
499 |
|> fold (curry op COMP) ex1s
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
500 |
|> Thm.implies_intr (ihyp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
501 |
|> Thm.implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
502 |
|> Thm.forall_intr (cterm_of thy x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
503 |
|> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
504 |
|> (fn it => fold (Thm.forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
505 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
506 |
val goalstate = Conjunction.intr graph_is_function complete
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
507 |
|> Thm.close_derivation
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
508 |
|> Goal.protect
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
509 |
|> fold_rev (Thm.implies_intr o cprop_of) compat
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
510 |
|> Thm.implies_intr (cprop_of complete)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
511 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
512 |
(goalstate, values)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
513 |
end
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
514 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
515 |
|
2651
|
516 |
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
517 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
518 |
(* wrapper -- restores quantifiers in rule specifications *)
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
519 |
fun inductive_def eqvt_flag (binding as ((R, T), _)) intrs lthy =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
520 |
let
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
521 |
val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, raw_induct, ...}, lthy) =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
522 |
lthy
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
523 |
|> Local_Theory.conceal
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
524 |
|> Inductive.add_inductive_i
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
525 |
{quiet_mode = true,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
526 |
verbose = ! trace,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
527 |
alt_name = Binding.empty,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
528 |
coind = false,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
529 |
no_elim = false,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
530 |
no_ind = false,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
531 |
skip_mono = true,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
532 |
fork_mono = false}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
533 |
[binding] (* relation *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
534 |
[] (* no parameters *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
535 |
(map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
536 |
[] (* no special monos *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
537 |
||> Local_Theory.restore_naming lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
538 |
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
539 |
val eqvt_thm' =
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
540 |
if eqvt_flag = false then NONE
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
541 |
else
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
542 |
let
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
543 |
val ([eqvt_thm], lthy) = Nominal_Eqvt.raw_equivariance false [Rdef] raw_induct intrs_gen lthy
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
544 |
in
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
545 |
SOME ((Nominal_ThmDecls.eqvt_transform lthy eqvt_thm) RS @{thm eqvtI})
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
546 |
end
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
547 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
548 |
val cert = cterm_of (ProofContext.theory_of lthy)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
549 |
fun requantify orig_intro thm =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
550 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
551 |
val (qs, t) = dest_all_all orig_intro
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
552 |
val frees = frees_in_term lthy t |> remove (op =) (Binding.name_of R, T)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
553 |
val vars = Term.add_vars (prop_of thm) [] |> rev
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
554 |
val varmap = AList.lookup (op =) (frees ~~ map fst vars)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
555 |
#> the_default ("",0)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
556 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
557 |
fold_rev (fn Free (n, T) =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
558 |
forall_intr_rename (n, cert (Var (varmap (n, T), T)))) qs thm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
559 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
560 |
in
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
561 |
((Rdef, map2 requantify intrs intrs_gen, forall_intr_vars elim_gen, induct, eqvt_thm'), lthy)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
562 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
563 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
564 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
565 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
566 |
fun define_graph Gname fvar domT ranT clauses RCss lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
567 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
568 |
val GT = domT --> ranT --> boolT
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
569 |
val (Gvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Gname, GT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
570 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
571 |
fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
572 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
573 |
fun mk_h_assm (rcfix, rcassm, rcarg) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
574 |
HOLogic.mk_Trueprop (Free Gvar $ rcarg $ (fvar $ rcarg))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
575 |
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
576 |
|> fold_rev (Logic.all o Free) rcfix
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
577 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
578 |
HOLogic.mk_Trueprop (Free Gvar $ lhs $ rhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
579 |
|> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
580 |
|> fold_rev (curry Logic.mk_implies) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
581 |
|> fold_rev Logic.all (fvar :: qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
582 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
583 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
584 |
val G_intros = map2 mk_GIntro clauses RCss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
585 |
in
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
586 |
inductive_def true ((Binding.name n, T), NoSyn) G_intros lthy
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
587 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
588 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
589 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
590 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
591 |
fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
592 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
593 |
val f_def =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
594 |
Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
595 |
$ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
596 |
|> Syntax.check_term lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
597 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
598 |
Local_Theory.define
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
599 |
((Binding.name (function_name fname), mixfix),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
600 |
((Binding.conceal (Binding.name fdefname), []), f_def)) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
601 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
602 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
603 |
fun define_recursion_relation Rname domT qglrs clauses RCss lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
604 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
605 |
val RT = domT --> domT --> boolT
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
606 |
val (Rvar as (n, T)) = singleton (Variable.variant_frees lthy []) (Rname, RT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
607 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
608 |
fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
609 |
HOLogic.mk_Trueprop (Free Rvar $ rcarg $ lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
610 |
|> fold_rev (curry Logic.mk_implies o prop_of) rcassm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
611 |
|> fold_rev (curry Logic.mk_implies) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
612 |
|> fold_rev (Logic.all o Free) rcfix
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
613 |
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
614 |
(* "!!qs xs. CS ==> G => (r, lhs) : R" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
615 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
616 |
val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
617 |
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
618 |
val ((R, RIntro_thms, R_elim, _, _), lthy) =
|
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
619 |
inductive_def false ((Binding.name n, T), NoSyn) (flat R_intross) lthy
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
620 |
in
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
621 |
((R, Library.unflat R_intross RIntro_thms, R_elim), lthy)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
622 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
623 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
624 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
625 |
fun fix_globals domT ranT fvar ctxt =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
626 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
627 |
val ([h, y, x, z, a, D, P, Pbool],ctxt') = Variable.variant_fixes
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
628 |
["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
629 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
630 |
(Globals {h = Free (h, domT --> ranT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
631 |
y = Free (y, ranT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
632 |
x = Free (x, domT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
633 |
z = Free (z, domT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
634 |
a = Free (a, domT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
635 |
D = Free (D, domT --> boolT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
636 |
P = Free (P, domT --> boolT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
637 |
Pbool = Free (Pbool, boolT),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
638 |
fvar = fvar,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
639 |
domT = domT,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
640 |
ranT = ranT},
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
641 |
ctxt')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
642 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
643 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
644 |
fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
645 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
646 |
fun inst_term t = subst_bound(f, abstract_over (fvar, t))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
647 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
648 |
(rcfix, map (Thm.assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
649 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
650 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
651 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
652 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
653 |
(**********************************************************
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
654 |
* PROVING THE RULES
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
655 |
**********************************************************)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
656 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
657 |
fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
658 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
659 |
val Globals {domT, z, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
660 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
661 |
fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
662 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
663 |
val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
664 |
val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
665 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
666 |
((Thm.assume z_smaller) RS ((Thm.assume lhs_acc) RS acc_downward))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
667 |
|> (fn it => it COMP graph_is_function)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
668 |
|> Thm.implies_intr z_smaller
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
669 |
|> Thm.forall_intr (cterm_of thy z)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
670 |
|> (fn it => it COMP valthm)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
671 |
|> Thm.implies_intr lhs_acc
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
672 |
|> asm_simplify (HOL_basic_ss addsimps [f_iff])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
673 |
|> fold_rev (Thm.implies_intr o cprop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
674 |
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
675 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
676 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
677 |
map2 mk_psimp clauses valthms
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
678 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
679 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
680 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
681 |
(** Induction rule **)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
682 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
683 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
684 |
val acc_subset_induct = @{thm predicate1I} RS @{thm accp_subset_induct}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
685 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
686 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
687 |
fun mk_partial_induct_rule thy globals R complete_thm clauses =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
688 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
689 |
val Globals {domT, x, z, a, P, D, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
690 |
val acc_R = mk_acc domT R
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
691 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
692 |
val x_D = Thm.assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
693 |
val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
694 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
695 |
val D_subset = cterm_of thy (Logic.all x
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
696 |
(Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
697 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
698 |
val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
699 |
Logic.all x (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
700 |
Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
701 |
HOLogic.mk_Trueprop (D $ z)))))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
702 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
703 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
704 |
(* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
705 |
val ihyp = Term.all domT $ Abs ("z", domT,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
706 |
Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
707 |
HOLogic.mk_Trueprop (P $ Bound 0)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
708 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
709 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
710 |
val aihyp = Thm.assume ihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
711 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
712 |
fun prove_case clause =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
713 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
714 |
val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...},
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
715 |
RCs, qglr = (oqs, _, _, _), ...} = clause
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
716 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
717 |
val case_hyp_conv = K (case_hyp RS eq_reflection)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
718 |
local open Conv in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
719 |
val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
720 |
val sih =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
721 |
fconv_rule (Conv.binder_conv
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
722 |
(K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
723 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
724 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
725 |
fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) = sih
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
726 |
|> Thm.forall_elim (cterm_of thy rcarg)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
727 |
|> Thm.elim_implies llRI
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
728 |
|> fold_rev (Thm.implies_intr o cprop_of) CCas
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
729 |
|> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
730 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
731 |
val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
732 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
733 |
val step = HOLogic.mk_Trueprop (P $ lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
734 |
|> fold_rev (curry Logic.mk_implies o prop_of) P_recs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
735 |
|> fold_rev (curry Logic.mk_implies) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
736 |
|> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
737 |
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
738 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
739 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
740 |
val P_lhs = Thm.assume step
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
741 |
|> fold Thm.forall_elim cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
742 |
|> Thm.elim_implies lhs_D
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
743 |
|> fold Thm.elim_implies ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
744 |
|> fold Thm.elim_implies P_recs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
745 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
746 |
val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
747 |
|> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
748 |
|> Thm.symmetric (* P lhs == P x *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
749 |
|> (fn eql => Thm.equal_elim eql P_lhs) (* "P x" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
750 |
|> Thm.implies_intr (cprop_of case_hyp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
751 |
|> fold_rev (Thm.implies_intr o cprop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
752 |
|> fold_rev Thm.forall_intr cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
753 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
754 |
(res, step)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
755 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
756 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
757 |
val (cases, steps) = split_list (map prove_case clauses)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
758 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
759 |
val istep = complete_thm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
760 |
|> Thm.forall_elim_vars 0
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
761 |
|> fold (curry op COMP) cases (* P x *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
762 |
|> Thm.implies_intr ihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
763 |
|> Thm.implies_intr (cprop_of x_D)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
764 |
|> Thm.forall_intr (cterm_of thy x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
765 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
766 |
val subset_induct_rule =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
767 |
acc_subset_induct
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
768 |
|> (curry op COMP) (Thm.assume D_subset)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
769 |
|> (curry op COMP) (Thm.assume D_dcl)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
770 |
|> (curry op COMP) (Thm.assume a_D)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
771 |
|> (curry op COMP) istep
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
772 |
|> fold_rev Thm.implies_intr steps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
773 |
|> Thm.implies_intr a_D
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
774 |
|> Thm.implies_intr D_dcl
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
775 |
|> Thm.implies_intr D_subset
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
776 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
777 |
val simple_induct_rule =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
778 |
subset_induct_rule
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
779 |
|> Thm.forall_intr (cterm_of thy D)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
780 |
|> Thm.forall_elim (cterm_of thy acc_R)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
781 |
|> assume_tac 1 |> Seq.hd
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
782 |
|> (curry op COMP) (acc_downward
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
783 |
|> (instantiate' [SOME (ctyp_of thy domT)]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
784 |
(map (SOME o cterm_of thy) [R, x, z]))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
785 |
|> Thm.forall_intr (cterm_of thy z)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
786 |
|> Thm.forall_intr (cterm_of thy x))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
787 |
|> Thm.forall_intr (cterm_of thy a)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
788 |
|> Thm.forall_intr (cterm_of thy P)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
789 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
790 |
simple_induct_rule
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
791 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
792 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
793 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
794 |
(* FIXME: broken by design *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
795 |
fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
796 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
797 |
val thy = ProofContext.theory_of ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
798 |
val ClauseInfo {cdata = ClauseContext {gs, lhs, cqs, ...},
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
799 |
qglr = (oqs, _, _, _), ...} = clause
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
800 |
val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
801 |
|> fold_rev (curry Logic.mk_implies) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
802 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
803 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
804 |
Goal.init goal
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
805 |
|> (SINGLE (resolve_tac [accI] 1)) |> the
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
806 |
|> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
807 |
|> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
808 |
|> Goal.conclude
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
809 |
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
810 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
811 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
812 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
813 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
814 |
(** Termination rule **)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
815 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
816 |
val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
817 |
val wf_in_rel = @{thm FunDef.wf_in_rel}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
818 |
val in_rel_def = @{thm FunDef.in_rel_def}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
819 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
820 |
fun mk_nest_term_case thy globals R' ihyp clause =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
821 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
822 |
val Globals {z, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
823 |
val ClauseInfo {cdata = ClauseContext {qs, cqs, ags, lhs, case_hyp, ...}, tree,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
824 |
qglr=(oqs, _, _, _), ...} = clause
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
825 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
826 |
val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
827 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
828 |
fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
829 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
830 |
val used = (u @ sub)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
831 |
|> map (fn (ctx,thm) => Function_Ctx_Tree.export_thm thy ctx thm)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
832 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
833 |
val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
834 |
|> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
835 |
|> Function_Ctx_Tree.export_term (fixes, assumes)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
836 |
|> fold_rev (curry Logic.mk_implies o prop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
837 |
|> fold_rev mk_forall_rename (map fst oqs ~~ qs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
838 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
839 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
840 |
val thm = Thm.assume hyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
841 |
|> fold Thm.forall_elim cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
842 |
|> fold Thm.elim_implies ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
843 |
|> Function_Ctx_Tree.import_thm thy (fixes, assumes)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
844 |
|> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
845 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
846 |
val z_eq_arg = HOLogic.mk_Trueprop (mk_eq (z, arg))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
847 |
|> cterm_of thy |> Thm.assume
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
848 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
849 |
val acc = thm COMP ih_case
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
850 |
val z_acc_local = acc
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
851 |
|> Conv.fconv_rule
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
852 |
(Conv.arg_conv (Conv.arg_conv (K (Thm.symmetric (z_eq_arg RS eq_reflection)))))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
853 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
854 |
val ethm = z_acc_local
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
855 |
|> Function_Ctx_Tree.export_thm thy (fixes,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
856 |
z_eq_arg :: case_hyp :: ags @ assumes)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
857 |
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
858 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
859 |
val sub' = sub @ [(([],[]), acc)]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
860 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
861 |
(sub', (hyp :: hyps, ethm :: thms))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
862 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
863 |
| step _ _ _ _ = raise Match
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
864 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
865 |
Function_Ctx_Tree.traverse_tree step tree
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
866 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
867 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
868 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
869 |
fun mk_nest_term_rule thy globals R R_cases clauses =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
870 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
871 |
val Globals { domT, x, z, ... } = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
872 |
val acc_R = mk_acc domT R
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
873 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
874 |
val R' = Free ("R", fastype_of R)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
875 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
876 |
val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
877 |
val inrel_R = Const (@{const_name FunDef.in_rel},
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
878 |
HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
879 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
880 |
val wfR' = HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP},
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
881 |
(domT --> domT --> boolT) --> boolT) $ R')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
882 |
|> cterm_of thy (* "wf R'" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
883 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
884 |
(* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
885 |
val ihyp = Term.all domT $ Abs ("z", domT,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
886 |
Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
887 |
HOLogic.mk_Trueprop (acc_R $ Bound 0)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
888 |
|> cterm_of thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
889 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
890 |
val ihyp_a = Thm.assume ihyp |> Thm.forall_elim_vars 0
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
891 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
892 |
val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
893 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
894 |
val (hyps, cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([], [])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
895 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
896 |
R_cases
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
897 |
|> Thm.forall_elim (cterm_of thy z)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
898 |
|> Thm.forall_elim (cterm_of thy x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
899 |
|> Thm.forall_elim (cterm_of thy (acc_R $ z))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
900 |
|> curry op COMP (Thm.assume R_z_x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
901 |
|> fold_rev (curry op COMP) cases
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
902 |
|> Thm.implies_intr R_z_x
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
903 |
|> Thm.forall_intr (cterm_of thy z)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
904 |
|> (fn it => it COMP accI)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
905 |
|> Thm.implies_intr ihyp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
906 |
|> Thm.forall_intr (cterm_of thy x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
907 |
|> (fn it => Drule.compose_single(it,2,wf_induct_rule))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
908 |
|> curry op RS (Thm.assume wfR')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
909 |
|> forall_intr_vars
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
910 |
|> (fn it => it COMP allI)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
911 |
|> fold Thm.implies_intr hyps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
912 |
|> Thm.implies_intr wfR'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
913 |
|> Thm.forall_intr (cterm_of thy R')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
914 |
|> Thm.forall_elim (cterm_of thy (inrel_R))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
915 |
|> curry op RS wf_in_rel
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
916 |
|> full_simplify (HOL_basic_ss addsimps [in_rel_def])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
917 |
|> Thm.forall_intr (cterm_of thy Rrel)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
918 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
919 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
920 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
921 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
922 |
(* Tail recursion (probably very fragile)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
923 |
*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
924 |
* FIXME:
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
925 |
* - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
926 |
* - Must we really replace the fvar by f here?
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
927 |
* - Splitting is not configured automatically: Problems with case?
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
928 |
*)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
929 |
fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
930 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
931 |
val Globals {domT, ranT, fvar, ...} = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
932 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
933 |
val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
934 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
935 |
val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
936 |
Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
937 |
(HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
938 |
(fn {prems=[a], ...} =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
939 |
((rtac (G_induct OF [a]))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
940 |
THEN_ALL_NEW rtac accI
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
941 |
THEN_ALL_NEW etac R_cases
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
942 |
THEN_ALL_NEW asm_full_simp_tac (simpset_of octxt)) 1)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
943 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
944 |
val default_thm =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
945 |
forall_intr_vars graph_implies_dom COMP (f_def COMP fundef_default_value)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
946 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
947 |
fun mk_trsimp clause psimp =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
948 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
949 |
val ClauseInfo {qglr = (oqs, _, _, _), cdata =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
950 |
ClauseContext {ctxt, cqs, gs, lhs, rhs, ...}, ...} = clause
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
951 |
val thy = ProofContext.theory_of ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
952 |
val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
953 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
954 |
val trsimp = Logic.list_implies(gs,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
955 |
HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
956 |
val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
957 |
fun simp_default_tac ss =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
958 |
asm_full_simp_tac (ss addsimps [default_thm, Let_def])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
959 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
960 |
Goal.prove ctxt [] [] trsimp (fn _ =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
961 |
rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
962 |
THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
963 |
THEN (simp_default_tac (simpset_of ctxt) 1)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
964 |
THEN TRY ((etac not_acc_down 1)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
965 |
THEN ((etac R_cases)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
966 |
THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
967 |
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
968 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
969 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
970 |
map2 mk_trsimp clauses psimps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
971 |
end
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
972 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
973 |
|
2650
e5fa8de0e4bd
derived equivariance for the function graph and function relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
974 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
975 |
fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
976 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
977 |
val FunctionConfig {domintros, tailrec, default=default_opt, ...} = config
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
978 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
979 |
val default_str = the_default "%x. undefined" default_opt (*FIXME dynamic scoping*)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
980 |
val fvar = Free (fname, fT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
981 |
val domT = domain_type fT
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
982 |
val ranT = range_type fT
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
983 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
984 |
val default = Syntax.parse_term lthy default_str
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
985 |
|> Type.constraint fT |> Syntax.check_term lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
986 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
987 |
val (globals, ctxt') = fix_globals domT ranT fvar lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
988 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
989 |
val Globals { x, h, ... } = globals
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
990 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
991 |
val clauses = map (mk_clause_context x ctxt') abstract_qglrs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
992 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
993 |
val n = length abstract_qglrs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
994 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
995 |
fun build_tree (ClauseContext { ctxt, rhs, ...}) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
996 |
Function_Ctx_Tree.mk_tree (fname, fT) h ctxt rhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
997 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
998 |
val trees = map build_tree clauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
999 |
val RCss = map find_calls trees
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1000 |
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1001 |
val ((G, GIntro_thms, G_elim, G_induct, SOME G_eqvt), lthy) =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1002 |
PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1003 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1004 |
val ((f, (_, f_defthm)), lthy) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1005 |
PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1006 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1007 |
val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1008 |
val trees = map (Function_Ctx_Tree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1009 |
|
2662
7c5bca978886
eliminated tracing code; added flag so that equivariance is only proved for the function graph, not the relation
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1010 |
val ((R, RIntro_thmss, R_elim), lthy) =
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1011 |
PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT abstract_qglrs clauses RCss) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1012 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1013 |
val (_, lthy) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1014 |
Local_Theory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1015 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1016 |
val newthy = ProofContext.theory_of lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1017 |
val clauses = map (transfer_clause_ctx newthy) clauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1018 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1019 |
val cert = cterm_of (ProofContext.theory_of lthy)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1020 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1021 |
val xclauses = PROFILE "xclauses"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1022 |
(map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1023 |
RCss GIntro_thms) RIntro_thmss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1024 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1025 |
val complete =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1026 |
mk_completeness globals clauses abstract_qglrs |> cert |> Thm.assume
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1027 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1028 |
val compat =
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1029 |
mk_compat_proof_obligations domT ranT fvar f abstract_qglrs
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1030 |
|> map (cert #> Thm.assume)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1031 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1032 |
val compat_store = store_compat_thms n compat
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1033 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1034 |
val (goalstate, values) = PROFILE "prove_stuff"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1035 |
(prove_stuff lthy globals G f R xclauses complete compat
|
2655
1c3ad1256f16
instantiated fundef_ex1_eqvt_at theorem with the indction hypothesis
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1036 |
compat_store G_elim G_eqvt) f_defthm
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1037 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1038 |
val mk_trsimps =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1039 |
mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1040 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1041 |
fun mk_partial_rules provedgoal =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1042 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1043 |
val newthy = theory_of_thm provedgoal (*FIXME*)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1044 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1045 |
val (graph_is_function, complete_thm) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1046 |
provedgoal
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1047 |
|> Conjunction.elim
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1048 |
|> apfst (Thm.forall_elim_vars 0)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1049 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1050 |
val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1051 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1052 |
val psimps = PROFILE "Proving simplification rules"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1053 |
(mk_psimps newthy globals R xclauses values f_iff) graph_is_function
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1054 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1055 |
val simple_pinduct = PROFILE "Proving partial induction rule"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1056 |
(mk_partial_induct_rule newthy globals R complete_thm) xclauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1057 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1058 |
val total_intro = PROFILE "Proving nested termination rule"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1059 |
(mk_nest_term_rule newthy globals R R_elim) xclauses
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1060 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1061 |
val dom_intros =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1062 |
if domintros then SOME (PROFILE "Proving domain introduction rules"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1063 |
(map (mk_domain_intro lthy globals R R_elim)) xclauses)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1064 |
else NONE
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1065 |
val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1066 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1067 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1068 |
FunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1069 |
psimps=psimps, simple_pinducts=[simple_pinduct],
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1070 |
termination=total_intro, trsimps=trsimps,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1071 |
domintros=dom_intros}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1072 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1073 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1074 |
((f, goalstate, mk_partial_rules), lthy)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1075 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1076 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1077 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1078 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1079 |
open Function_Lib
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1080 |
open Function_Common
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1081 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1082 |
type qgar = string * (string * typ) list * term list * term list * term
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1083 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1084 |
datatype mutual_part = MutualPart of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1085 |
{i : int,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1086 |
i' : int,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1087 |
fvar : string * typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1088 |
cargTs: typ list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1089 |
f_def: term,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1090 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1091 |
f: term option,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1092 |
f_defthm : thm option}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1093 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1094 |
datatype mutual_info = Mutual of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1095 |
{n : int,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1096 |
n' : int,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1097 |
fsum_var : string * typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1098 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1099 |
ST: typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1100 |
RST: typ,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1101 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1102 |
parts: mutual_part list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1103 |
fqgars: qgar list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1104 |
qglrs: ((string * typ) list * term list * term * term) list,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1105 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1106 |
fsum : term option}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1107 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1108 |
fun mutual_induct_Pnames n =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1109 |
if n < 5 then fst (chop n ["P","Q","R","S"])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1110 |
else map (fn i => "P" ^ string_of_int i) (1 upto n)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1111 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1112 |
fun get_part fname =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1113 |
the o find_first (fn (MutualPart {fvar=(n,_), ...}) => n = fname)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1114 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1115 |
(* FIXME *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1116 |
fun mk_prod_abs e (t1, t2) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1117 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1118 |
val bTs = rev (map snd e)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1119 |
val T1 = fastype_of1 (bTs, t1)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1120 |
val T2 = fastype_of1 (bTs, t2)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1121 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1122 |
HOLogic.pair_const T1 T2 $ t1 $ t2
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1123 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1124 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1125 |
fun analyze_eqs ctxt defname fs eqs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1126 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1127 |
val num = length fs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1128 |
val fqgars = map (split_def ctxt (K true)) eqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1129 |
val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1130 |
|> AList.lookup (op =) #> the
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1131 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1132 |
fun curried_types (fname, fT) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1133 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1134 |
val (caTs, uaTs) = chop (arity_of fname) (binder_types fT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1135 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1136 |
(caTs, uaTs ---> body_type fT)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1137 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1138 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1139 |
val (caTss, resultTs) = split_list (map curried_types fs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1140 |
val argTs = map (foldr1 HOLogic.mk_prodT) caTss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1141 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1142 |
val dresultTs = distinct (op =) resultTs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1143 |
val n' = length dresultTs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1144 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1145 |
val RST = Balanced_Tree.make (uncurry SumTree.mk_sumT) dresultTs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1146 |
val ST = Balanced_Tree.make (uncurry SumTree.mk_sumT) argTs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1147 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1148 |
val fsum_type = ST --> RST
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1149 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1150 |
val ([fsum_var_name], _) = Variable.add_fixes [ defname ^ "_sum" ] ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1151 |
val fsum_var = (fsum_var_name, fsum_type)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1152 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1153 |
fun define (fvar as (n, _)) caTs resultT i =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1154 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1155 |
val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1156 |
val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1157 |
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1158 |
val f_exp = SumTree.mk_proj RST n' i'
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1159 |
(Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars))
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1160 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1161 |
val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1162 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1163 |
val rew = (n, fold_rev lambda vars f_exp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1164 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1165 |
(MutualPart {i=i, i'=i', fvar=fvar,cargTs=caTs,f_def=def,f=NONE,f_defthm=NONE}, rew)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1166 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1167 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1168 |
val (parts, rews) = split_list (map4 define fs caTss resultTs (1 upto num))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1169 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1170 |
fun convert_eqs (f, qs, gs, args, rhs) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1171 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1172 |
val MutualPart {i, i', ...} = get_part f parts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1173 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1174 |
(qs, gs, SumTree.mk_inj ST num i (foldr1 (mk_prod_abs qs) args),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1175 |
SumTree.mk_inj RST n' i' (replace_frees rews rhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1176 |
|> Envir.beta_norm)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1177 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1178 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1179 |
val qglrs = map convert_eqs fqgars
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1180 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1181 |
Mutual {n=num, n'=n', fsum_var=fsum_var, ST=ST, RST=RST,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1182 |
parts=parts, fqgars=fqgars, qglrs=qglrs, fsum=NONE}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1183 |
end
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1184 |
*}
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1185 |
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1186 |
ML {*
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1187 |
fun define_projections fixes mutual fsum lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1188 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1189 |
fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1190 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1191 |
val ((f, (_, f_defthm)), lthy') =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1192 |
Local_Theory.define
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1193 |
((Binding.name fname, mixfix),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1194 |
((Binding.conceal (Binding.name (fname ^ "_def")), []),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1195 |
Term.subst_bound (fsum, f_def))) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1196 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1197 |
(MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs=cargTs, f_def=f_def,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1198 |
f=SOME f, f_defthm=SOME f_defthm },
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1199 |
lthy')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1200 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1201 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1202 |
val Mutual { n, n', fsum_var, ST, RST, parts, fqgars, qglrs, ... } = mutual
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1203 |
val (parts', lthy') = fold_map def (parts ~~ fixes) lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1204 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1205 |
(Mutual { n=n, n'=n', fsum_var=fsum_var, ST=ST, RST=RST, parts=parts',
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1206 |
fqgars=fqgars, qglrs=qglrs, fsum=SOME fsum },
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1207 |
lthy')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1208 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1209 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1210 |
fun in_context ctxt (f, pre_qs, pre_gs, pre_args, pre_rhs) F =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1211 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1212 |
val thy = ProofContext.theory_of ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1213 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1214 |
val oqnames = map fst pre_qs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1215 |
val (qs, _) = Variable.variant_fixes oqnames ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1216 |
|>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1217 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1218 |
fun inst t = subst_bounds (rev qs, t)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1219 |
val gs = map inst pre_gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1220 |
val args = map inst pre_args
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1221 |
val rhs = inst pre_rhs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1222 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1223 |
val cqs = map (cterm_of thy) qs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1224 |
val ags = map (Thm.assume o cterm_of thy) gs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1225 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1226 |
val import = fold Thm.forall_elim cqs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1227 |
#> fold Thm.elim_implies ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1228 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1229 |
val export = fold_rev (Thm.implies_intr o cprop_of) ags
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1230 |
#> fold_rev forall_intr_rename (oqnames ~~ cqs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1231 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1232 |
F ctxt (f, qs, gs, args, rhs) import export
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1233 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1234 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1235 |
fun recover_mutual_psimp all_orig_fdefs parts ctxt (fname, _, _, args, rhs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1236 |
import (export : thm -> thm) sum_psimp_eq =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1237 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1238 |
val (MutualPart {f=SOME f, ...}) = get_part fname parts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1239 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1240 |
val psimp = import sum_psimp_eq
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1241 |
val (simp, restore_cond) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1242 |
case cprems_of psimp of
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1243 |
[] => (psimp, I)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1244 |
| [cond] => (Thm.implies_elim psimp (Thm.assume cond), Thm.implies_intr cond)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1245 |
| _ => raise General.Fail "Too many conditions"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1246 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1247 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1248 |
Goal.prove ctxt [] []
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1249 |
(HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1250 |
(fn _ => (Local_Defs.unfold_tac ctxt all_orig_fdefs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1251 |
THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1252 |
THEN (simp_tac (simpset_of ctxt)) 1) (* FIXME: global simpset?!! *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1253 |
|> restore_cond
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1254 |
|> export
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1255 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1256 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1257 |
fun mk_applied_form ctxt caTs thm =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1258 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1259 |
val thy = ProofContext.theory_of ctxt
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1260 |
val xs = map_index (fn (i,T) => cterm_of thy (Free ("x" ^ string_of_int i, T))) caTs (* FIXME: Bind xs properly *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1261 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1262 |
fold (fn x => fn thm => Thm.combination thm (Thm.reflexive x)) xs thm
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1263 |
|> Conv.fconv_rule (Thm.beta_conversion true)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1264 |
|> fold_rev Thm.forall_intr xs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1265 |
|> Thm.forall_elim_vars 0
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1266 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1267 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1268 |
fun mutual_induct_rules lthy induct all_f_defs (Mutual {n, ST, parts, ...}) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1269 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1270 |
val cert = cterm_of (ProofContext.theory_of lthy)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1271 |
val newPs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1272 |
map2 (fn Pname => fn MutualPart {cargTs, ...} =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1273 |
Free (Pname, cargTs ---> HOLogic.boolT))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1274 |
(mutual_induct_Pnames (length parts)) parts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1275 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1276 |
fun mk_P (MutualPart {cargTs, ...}) P =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1277 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1278 |
val avars = map_index (fn (i,T) => Var (("a", i), T)) cargTs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1279 |
val atup = foldr1 HOLogic.mk_prod avars
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1280 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1281 |
HOLogic.tupled_lambda atup (list_comb (P, avars))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1282 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1283 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1284 |
val Ps = map2 mk_P parts newPs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1285 |
val case_exp = SumTree.mk_sumcases HOLogic.boolT Ps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1286 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1287 |
val induct_inst =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1288 |
Thm.forall_elim (cert case_exp) induct
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1289 |
|> full_simplify SumTree.sumcase_split_ss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1290 |
|> full_simplify (HOL_basic_ss addsimps all_f_defs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1291 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1292 |
fun project rule (MutualPart {cargTs, i, ...}) k =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1293 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1294 |
val afs = map_index (fn (j,T) => Free ("a" ^ string_of_int (j + k), T)) cargTs (* FIXME! *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1295 |
val inj = SumTree.mk_inj ST n i (foldr1 HOLogic.mk_prod afs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1296 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1297 |
(rule
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1298 |
|> Thm.forall_elim (cert inj)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1299 |
|> full_simplify SumTree.sumcase_split_ss
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1300 |
|> fold_rev (Thm.forall_intr o cert) (afs @ newPs),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1301 |
k + length cargTs)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1302 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1303 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1304 |
fst (fold_map (project induct_inst) parts 0)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1305 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1306 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1307 |
fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1308 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1309 |
val result = inner_cont proof
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1310 |
val FunctionResult {G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1311 |
termination, domintros, ...} = result
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1312 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1313 |
val (all_f_defs, fs) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1314 |
map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1315 |
(mk_applied_form lthy cargTs (Thm.symmetric f_def), f))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1316 |
parts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1317 |
|> split_list
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1318 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1319 |
val all_orig_fdefs =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1320 |
map (fn MutualPart {f_defthm = SOME f_def, ...} => f_def) parts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1321 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1322 |
fun mk_mpsimp fqgar sum_psimp =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1323 |
in_context lthy fqgar (recover_mutual_psimp all_orig_fdefs parts) sum_psimp
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1324 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1325 |
val rew_ss = HOL_basic_ss addsimps all_f_defs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1326 |
val mpsimps = map2 mk_mpsimp fqgars psimps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1327 |
val mtrsimps = Option.map (map2 mk_mpsimp fqgars) trsimps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1328 |
val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1329 |
val mtermination = full_simplify rew_ss termination
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1330 |
val mdomintros = Option.map (map (full_simplify rew_ss)) domintros
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1331 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1332 |
FunctionResult { fs=fs, G=G, R=R,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1333 |
psimps=mpsimps, simple_pinducts=minducts,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1334 |
cases=cases, termination=mtermination,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1335 |
domintros=mdomintros, trsimps=mtrsimps}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1336 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1337 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1338 |
fun prepare_function_mutual config defname fixes eqss lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1339 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1340 |
val mutual as Mutual {fsum_var=(n, T), qglrs, ...} =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1341 |
analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1342 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1343 |
val ((fsum, goalstate, cont), lthy') =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1344 |
prepare_function config defname [((n, T), NoSyn)] qglrs lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1345 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1346 |
val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1347 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1348 |
val mutual_cont = mk_partial_rules_mutual lthy'' cont mutual'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1349 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1350 |
((goalstate, mutual_cont), lthy'')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1351 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1352 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1353 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1354 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1355 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1356 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1357 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1358 |
open Function_Lib
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1359 |
open Function_Common
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1360 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1361 |
val simp_attribs = map (Attrib.internal o K)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1362 |
[Simplifier.simp_add,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1363 |
Code.add_default_eqn_attribute,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1364 |
Nitpick_Simps.add]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1365 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1366 |
val psimp_attribs = map (Attrib.internal o K)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1367 |
[Nitpick_Psimps.add]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1368 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1369 |
fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1370 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1371 |
fun add_simps fnames post sort extra_qualify label mod_binding moreatts
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1372 |
simps lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1373 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1374 |
val spec = post simps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1375 |
|> map (apfst (apsnd (fn ats => moreatts @ ats)))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1376 |
|> map (apfst (apfst extra_qualify))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1377 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1378 |
val (saved_spec_simps, lthy) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1379 |
fold_map Local_Theory.note spec lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1380 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1381 |
val saved_simps = maps snd saved_spec_simps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1382 |
val simps_by_f = sort saved_simps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1383 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1384 |
fun add_for_f fname simps =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1385 |
Local_Theory.note
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1386 |
((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1387 |
#> snd
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1388 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1389 |
(saved_simps, fold2 add_for_f fnames simps_by_f lthy)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1390 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1391 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1392 |
fun prepare_function is_external prep default_constraint fixspec eqns config lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1393 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1394 |
val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1395 |
val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1396 |
val fixes = map (apfst (apfst Binding.name_of)) fixes0;
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1397 |
val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1398 |
val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1399 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1400 |
val defname = mk_defname fixes
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1401 |
val FunctionConfig {partials, tailrec, default, ...} = config
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1402 |
val _ =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1403 |
if tailrec then Output.legacy_feature
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1404 |
"'function (tailrec)' command. Use 'partial_function (tailrec)'."
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1405 |
else ()
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1406 |
val _ =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1407 |
if is_some default then Output.legacy_feature
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1408 |
"'function (default)'. Use 'partial_function'."
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1409 |
else ()
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1410 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1411 |
val ((goal_state, cont), lthy') =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1412 |
prepare_function_mutual config defname fixes eqs lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1413 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1414 |
fun afterqed [[proof]] lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1415 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1416 |
val FunctionResult {fs, R, psimps, trsimps, simple_pinducts,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1417 |
termination, domintros, cases, ...} =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1418 |
cont (Thm.close_derivation proof)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1419 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1420 |
val fnames = map (fst o fst) fixes
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1421 |
fun qualify n = Binding.name n
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1422 |
|> Binding.qualify true defname
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1423 |
val conceal_partial = if partials then I else Binding.conceal
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1424 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1425 |
val addsmps = add_simps fnames post sort_cont
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1426 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1427 |
val (((psimps', pinducts'), (_, [termination'])), lthy) =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1428 |
lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1429 |
|> addsmps (conceal_partial o Binding.qualify false "partial")
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1430 |
"psimps" conceal_partial psimp_attribs psimps
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1431 |
||> (case trsimps of NONE => I | SOME thms =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1432 |
addsmps I "simps" I simp_attribs thms #> snd
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1433 |
#> Spec_Rules.add Spec_Rules.Equational (fs, thms))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1434 |
||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1435 |
[Attrib.internal (K (Rule_Cases.case_names cnames)),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1436 |
Attrib.internal (K (Rule_Cases.consumes 1)),
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1437 |
Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1438 |
||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1439 |
||> (snd o Local_Theory.note ((qualify "cases",
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1440 |
[Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1441 |
||> (case domintros of NONE => I | SOME thms =>
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1442 |
Local_Theory.note ((qualify "domintros", []), thms) #> snd)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1443 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1444 |
val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1445 |
pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1446 |
fs=fs, R=R, defname=defname, is_partial=true }
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1447 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1448 |
val _ =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1449 |
if not is_external then ()
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1450 |
else Specification.print_consts lthy (K false) (map fst fixes)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1451 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1452 |
(info,
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1453 |
lthy |> Local_Theory.declaration false (add_function_data o morph_function_data info))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1454 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1455 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1456 |
((goal_state, afterqed), lthy')
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1457 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1458 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1459 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1460 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1461 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1462 |
fun gen_function is_external prep default_constraint fixspec eqns config lthy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1463 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1464 |
val ((goal_state, afterqed), lthy') =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1465 |
prepare_function is_external prep default_constraint fixspec eqns config lthy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1466 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1467 |
lthy'
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1468 |
|> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (concl_of goal_state), [])]]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1469 |
|> Proof.refine (Method.primitive_text (K goal_state))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1470 |
|> Seq.hd
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1471 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1472 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1473 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1474 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1475 |
ML {*
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1476 |
val function = gen_function false Specification.check_spec (Type_Infer.anyT HOLogic.typeS)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1477 |
val function_cmd = gen_function true Specification.read_spec "_::type"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1478 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1479 |
fun add_case_cong n thy =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1480 |
let
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1481 |
val cong = #case_cong (Datatype.the_info thy n)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1482 |
|> safe_mk_meta_eq
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1483 |
in
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1484 |
Context.theory_map
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1485 |
(Function_Ctx_Tree.map_function_congs (Thm.add_thm cong)) thy
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1486 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1487 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1488 |
val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1489 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1490 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1491 |
(* setup *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1492 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1493 |
val setup =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1494 |
Attrib.setup @{binding fundef_cong}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1495 |
(Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1496 |
"declaration of congruence rule for function definitions"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1497 |
#> setup_case_cong
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1498 |
#> Function_Relation.setup
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1499 |
#> Function_Common.Termination_Simps.setup
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1500 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1501 |
val get_congs = Function_Ctx_Tree.get_function_congs
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1502 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1503 |
fun get_info ctxt t = Item_Net.retrieve (get_function ctxt) t
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1504 |
|> the_single |> snd
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1505 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1506 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1507 |
(* outer syntax *)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1508 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1509 |
val _ =
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1510 |
Outer_Syntax.local_theory_to_proof "nominal_primrec" "define recursive functions for nominal types"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1511 |
Keyword.thy_goal
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1512 |
(function_parser default_config
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1513 |
>> (fn ((config, fixes), statements) => function_cmd fixes statements config))
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1514 |
*}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1515 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1516 |
ML {* trace := true *}
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1517 |
|
2654
|
1518 |
lemma test:
|
|
1519 |
assumes a: "[[x]]lst. t = [[x]]lst. t'"
|
|
1520 |
shows "t = t'"
|
|
1521 |
using a
|
|
1522 |
apply(simp add: Abs_eq_iff)
|
|
1523 |
apply(erule exE)
|
|
1524 |
apply(simp only: alphas)
|
|
1525 |
apply(erule conjE)+
|
|
1526 |
apply(rule sym)
|
|
1527 |
apply(clarify)
|
|
1528 |
apply(rule supp_perm_eq)
|
|
1529 |
apply(subgoal_tac "set [x] \<sharp>* p")
|
|
1530 |
apply(auto simp add: fresh_star_def)[1]
|
|
1531 |
apply(simp)
|
|
1532 |
apply(simp add: fresh_star_def)
|
|
1533 |
apply(simp add: fresh_perm)
|
|
1534 |
done
|
|
1535 |
|
|
1536 |
lemma test2:
|
|
1537 |
assumes "a \<sharp> x" "c \<sharp> x" "b \<sharp> y" "c \<sharp> y"
|
|
1538 |
and "(a \<rightleftharpoons> c) \<bullet> x = (b \<rightleftharpoons> c) \<bullet> y"
|
|
1539 |
shows "x = y"
|
|
1540 |
using assms by (simp add: swap_fresh_fresh)
|
|
1541 |
|
|
1542 |
lemma test3:
|
|
1543 |
assumes "x = y"
|
|
1544 |
and "a \<sharp> x" "c \<sharp> x" "b \<sharp> y" "c \<sharp> y"
|
|
1545 |
shows "(a \<rightleftharpoons> c) \<bullet> x = (b \<rightleftharpoons> c) \<bullet> y"
|
|
1546 |
using assms by (simp add: swap_fresh_fresh)
|
|
1547 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1548 |
nominal_primrec
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1549 |
depth :: "lam \<Rightarrow> nat"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1550 |
where
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1551 |
"depth (Var x) = 1"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1552 |
| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1553 |
| "depth (Lam x t) = (depth t) + 1"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1554 |
apply(rule_tac y="x" in lam.exhaust)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1555 |
apply(simp_all)[3]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1556 |
apply(simp_all only: lam.distinct)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1557 |
apply(simp add: lam.eq_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1558 |
apply(simp add: lam.eq_iff)
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1559 |
apply(subst (asm) Abs_eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1560 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1561 |
apply(simp add: alphas)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1562 |
apply(clarify)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1563 |
oops
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1564 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1565 |
lemma removeAll_eqvt[eqvt]:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1566 |
shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1567 |
by (induct xs) (auto)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1568 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1569 |
nominal_primrec
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1570 |
frees_lst :: "lam \<Rightarrow> atom list"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1571 |
where
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1572 |
"frees_lst (Var x) = [atom x]"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1573 |
| "frees_lst (App t1 t2) = (frees_lst t1) @ (frees_lst t2)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1574 |
| "frees_lst (Lam x t) = removeAll (atom x) (frees_lst t)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1575 |
apply(rule_tac y="x" in lam.exhaust)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1576 |
apply(simp_all)[3]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1577 |
apply(simp_all only: lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1578 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1579 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1580 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1581 |
apply(simp add: Abs_eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1582 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1583 |
apply(simp add: alphas)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1584 |
apply(simp add: atom_eqvt)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1585 |
apply(clarify)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1586 |
oops
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1587 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1588 |
nominal_primrec
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1589 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [100,100,100] 100)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1590 |
where
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1591 |
"(Var x)[y ::= s] = (if x=y then s else (Var x))"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1592 |
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1593 |
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1594 |
apply(case_tac x)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1595 |
apply(simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1596 |
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1597 |
apply(simp add: lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1598 |
apply(auto)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1599 |
apply(simp add: lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1600 |
apply(auto)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1601 |
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1602 |
apply(simp_all add: lam.distinct)[5]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1603 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1604 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1605 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1606 |
apply(erule conjE)+
|
2661
|
1607 |
oops
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1608 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1609 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1610 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1611 |
nominal_primrec
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1612 |
depth :: "lam \<Rightarrow> nat"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1613 |
where
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1614 |
"depth (Var x) = 1"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1615 |
| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1616 |
| "depth (Lam x t) = (depth t) + 1"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1617 |
apply(rule_tac y="x" in lam.exhaust)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1618 |
apply(simp_all)[3]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1619 |
apply(simp_all only: lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1620 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1621 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1622 |
(*
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1623 |
apply(subst (asm) Abs_eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1624 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1625 |
apply(simp add: alphas)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1626 |
apply(clarify)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1627 |
*)
|
2654
|
1628 |
apply(subgoal_tac "finite (supp (x, xa, t, ta, depth_sumC t, depth_sumC ta))")
|
|
1629 |
apply(erule_tac ?'a="name" in obtain_at_base)
|
|
1630 |
unfolding fresh_def[symmetric]
|
|
1631 |
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
|
|
1632 |
apply(simp add: Abs_fresh_iff)
|
|
1633 |
apply(simp add: Abs_fresh_iff)
|
|
1634 |
apply(simp add: Abs_fresh_iff)
|
|
1635 |
apply(simp add: Abs_fresh_iff)
|
|
1636 |
apply(rule_tac a="atom x" and b="atom xa" and c="atom a" in test2)
|
|
1637 |
apply(simp add: pure_fresh)
|
|
1638 |
apply(simp add: pure_fresh)
|
|
1639 |
apply(simp add: pure_fresh)
|
|
1640 |
apply(simp add: pure_fresh)
|
|
1641 |
apply(simp add: eqvt_at_def)
|
|
1642 |
apply(drule test)
|
|
1643 |
apply(simp)
|
|
1644 |
apply(simp add: finite_supp)
|
|
1645 |
done
|
|
1646 |
|
|
1647 |
termination depth
|
|
1648 |
apply(relation "measure size")
|
|
1649 |
apply(auto simp add: lam.size)
|
|
1650 |
done
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1651 |
|
2651
|
1652 |
thm depth.psimps
|
2654
|
1653 |
thm depth.simps
|
2651
|
1654 |
|
2657
|
1655 |
|
|
1656 |
lemma swap_set_not_in_at:
|
|
1657 |
fixes a b::"'a::at"
|
|
1658 |
and S::"'a::at set"
|
|
1659 |
assumes a: "a \<notin> S" "b \<notin> S"
|
|
1660 |
shows "(a \<leftrightarrow> b) \<bullet> S = S"
|
|
1661 |
unfolding permute_set_eq
|
|
1662 |
using a by (auto simp add: permute_flip_at)
|
|
1663 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1664 |
lemma removeAll_eqvt[eqvt]:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1665 |
shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1666 |
by (induct xs) (auto)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1667 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1668 |
nominal_primrec
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1669 |
frees_lst :: "lam \<Rightarrow> atom list"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1670 |
where
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1671 |
"frees_lst (Var x) = [atom x]"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1672 |
| "frees_lst (App t1 t2) = (frees_lst t1) @ (frees_lst t2)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1673 |
| "frees_lst (Lam x t) = removeAll (atom x) (frees_lst t)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1674 |
apply(rule_tac y="x" in lam.exhaust)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1675 |
apply(simp_all)[3]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1676 |
apply(simp_all only: lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1677 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1678 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1679 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1680 |
apply(simp add: Abs_eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1681 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1682 |
apply(simp add: alphas)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1683 |
apply(simp add: atom_eqvt)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1684 |
apply(clarify)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1685 |
apply(rule trans)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1686 |
apply(rule sym)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1687 |
apply(rule_tac p="p" in supp_perm_eq)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1688 |
oops
|
2657
|
1689 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1690 |
nominal_primrec
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1691 |
frees :: "lam \<Rightarrow> name set"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1692 |
where
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1693 |
"frees (Var x) = {x}"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1694 |
| "frees (App t1 t2) = (frees t1) \<union> (frees t2)"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1695 |
| "frees (Lam x t) = (frees t) - {x}"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1696 |
apply(rule_tac y="x" in lam.exhaust)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1697 |
apply(simp_all)[3]
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1698 |
apply(simp_all only: lam.distinct)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1699 |
apply(simp add: lam.eq_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1700 |
apply(simp add: lam.eq_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1701 |
apply(simp add: lam.eq_iff)
|
2654
|
1702 |
apply(subgoal_tac "finite (supp (x, xa, t, ta, frees_sumC t, frees_sumC ta))")
|
|
1703 |
apply(erule_tac ?'a="name" in obtain_at_base)
|
|
1704 |
unfolding fresh_def[symmetric]
|
|
1705 |
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
|
|
1706 |
apply(simp add: Abs_fresh_iff)
|
|
1707 |
apply(simp add: Abs_fresh_iff)
|
|
1708 |
apply(simp add: Abs_fresh_iff)
|
|
1709 |
apply(simp add: Abs_fresh_iff)
|
|
1710 |
apply(simp)
|
|
1711 |
apply(drule test)
|
2657
|
1712 |
apply(rule_tac t="frees_sumC t - {x}" and s="(x \<leftrightarrow> a) \<bullet> (frees_sumC t - {x})" in subst)
|
2654
|
1713 |
oops
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1714 |
|
2660
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1715 |
thm Abs_eq_iff[simplified alphas]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1716 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1717 |
lemma Abs_set_eq_iff2:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1718 |
fixes x y::"'a::fs"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1719 |
shows "[bs]set. x = [cs]set. y \<longleftrightarrow>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1720 |
(\<exists>p. supp ([bs]set. x) = supp ([cs]set. y) \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1721 |
supp ([bs]set. x) \<sharp>* p \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1722 |
p \<bullet> x = y \<and> p \<bullet> bs = cs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1723 |
unfolding Abs_eq_iff
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1724 |
unfolding alphas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1725 |
unfolding supp_Abs
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1726 |
by simp
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1727 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1728 |
lemma Abs_set_eq_iff3:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1729 |
fixes x y::"'a::fs"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1730 |
assumes a: "finite bs" "finite cs"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1731 |
shows "[bs]set. x = [cs]set. y \<longleftrightarrow>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1732 |
(\<exists>p. supp ([bs]set. x) = supp ([cs]set. y) \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1733 |
supp ([bs]set. x) \<sharp>* p \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1734 |
p \<bullet> x = y \<and> p \<bullet> bs = cs \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1735 |
supp p \<subseteq> bs \<union> cs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1736 |
unfolding Abs_eq_iff
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1737 |
unfolding alphas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1738 |
unfolding supp_Abs
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1739 |
apply(auto)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1740 |
using set_renaming_perm
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1741 |
sorry
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1742 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1743 |
lemma Abs_eq_iff2:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1744 |
fixes x y::"'a::fs"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1745 |
shows "[bs]lst. x = [cs]lst. y \<longleftrightarrow>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1746 |
(\<exists>p. supp ([bs]lst. x) = supp ([cs]lst. y) \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1747 |
supp ([bs]lst. x) \<sharp>* p \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1748 |
p \<bullet> x = y \<and> p \<bullet> bs = cs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1749 |
unfolding Abs_eq_iff
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1750 |
unfolding alphas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1751 |
unfolding supp_Abs
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1752 |
by simp
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1753 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1754 |
lemma Abs_eq_iff3:
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1755 |
fixes x y::"'a::fs"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1756 |
shows "[bs]lst. x = [cs]lst. y \<longleftrightarrow>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1757 |
(\<exists>p. supp ([bs]lst. x) = supp ([cs]lst. y) \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1758 |
supp ([bs]lst. x) \<sharp>* p \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1759 |
p \<bullet> x = y \<and> p \<bullet> bs = cs \<and>
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1760 |
supp p \<subseteq> set bs \<union> set cs)"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1761 |
unfolding Abs_eq_iff
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1762 |
unfolding alphas
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1763 |
unfolding supp_Abs
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1764 |
apply(auto)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1765 |
using list_renaming_perm
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1766 |
apply -
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1767 |
apply(drule_tac x="bs" in meta_spec)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1768 |
apply(drule_tac x="p" in meta_spec)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1769 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1770 |
apply(rule_tac x="q" in exI)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1771 |
apply(simp add: set_eqvt)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1772 |
sorry
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1773 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1774 |
nominal_primrec
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1775 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [100,100,100] 100)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1776 |
where
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1777 |
"(Var x)[y ::= s] = (if x=y then s else (Var x))"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1778 |
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1779 |
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1780 |
apply(case_tac x)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1781 |
apply(simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1782 |
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1783 |
apply(simp add: lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1784 |
apply(auto)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1785 |
apply(simp add: lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1786 |
apply(auto)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1787 |
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1788 |
apply(simp_all add: lam.distinct)[5]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1789 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1790 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1791 |
apply(simp add: lam.eq_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1792 |
apply(erule conjE)+
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1793 |
apply(subst (asm) Abs_eq_iff3)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1794 |
apply(erule exE)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1795 |
apply(erule conjE)+
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1796 |
apply(clarify)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1797 |
apply(perm_simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1798 |
apply(simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1799 |
apply(rule trans)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1800 |
apply(rule sym)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1801 |
apply(rule_tac p="p" in supp_perm_eq)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1802 |
apply(rule fresh_star_supp_conv)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1803 |
apply(drule fresh_star_supp_conv)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1804 |
apply(simp add: Abs_fresh_star_iff)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1805 |
thm fresh_eqvt_at
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1806 |
apply(rule_tac f="subst_sumC" in fresh_eqvt_at)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1807 |
apply(assumption)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1808 |
apply(simp add: finite_supp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1809 |
prefer 2
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1810 |
apply(simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1811 |
apply(simp add: eqvt_at_def)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1812 |
apply(perm_simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1813 |
apply(subgoal_tac "p \<bullet> ya = ya")
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1814 |
apply(subgoal_tac "p \<bullet> sa = sa")
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1815 |
apply(simp)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1816 |
apply(rule supp_perm_eq)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1817 |
apply(rule fresh_star_supp_conv)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1818 |
apply(auto simp add: fresh_star_def fresh_Pair)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1819 |
apply(rule supp_perm_eq)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1820 |
apply(rule fresh_star_supp_conv)
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1821 |
apply(auto simp add: fresh_star_def fresh_Pair)[1]
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1822 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1823 |
|
3342a2d13d95
nominal_function works now completely for frees and depth; still a propbelm with subst; no unproved assumptions
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1824 |
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1825 |
nominal_primrec
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1826 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [100,100,100] 100)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1827 |
where
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1828 |
"(Var x)[y ::= s] = (if x=y then s else (Var x))"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1829 |
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1830 |
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1831 |
apply(case_tac x)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1832 |
apply(simp)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1833 |
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1834 |
apply(simp add: lam.eq_iff lam.distinct)
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1835 |
apply(auto)[1]
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1836 |
apply(simp add: lam.eq_iff lam.distinct)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1837 |
apply(auto)[1]
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1838 |
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
|
2654
|
1839 |
apply(simp_all add: lam.distinct)[5]
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1840 |
apply(simp add: lam.eq_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1841 |
apply(simp add: lam.eq_iff)
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1842 |
apply(simp add: lam.eq_iff)
|
2654
|
1843 |
apply(subgoal_tac "finite (supp (ya, sa, x, xa, t, ta, subst_sumC (t, ya, sa), subst_sumC (ta, ya, sa)))")
|
|
1844 |
apply(erule_tac ?'a="name" in obtain_at_base)
|
|
1845 |
unfolding fresh_def[symmetric]
|
|
1846 |
apply(rule_tac a="atom x" and b="atom xa" and c="atom a" in test2)
|
|
1847 |
apply(simp add: Abs_fresh_iff)
|
|
1848 |
apply(simp add: Abs_fresh_iff)
|
|
1849 |
apply(simp add: Abs_fresh_iff)
|
|
1850 |
apply(simp add: Abs_fresh_iff)
|
|
1851 |
apply(erule conjE)+
|
|
1852 |
apply(drule_tac a="atom x" and b="atom xa" and c="atom a" in test3)
|
|
1853 |
apply(simp add: Abs_fresh_iff)
|
|
1854 |
apply(simp add: Abs_fresh_iff)
|
|
1855 |
apply(simp add: Abs_fresh_iff)
|
|
1856 |
apply(simp add: Abs_fresh_iff)
|
|
1857 |
apply(simp add: eqvt_at_def)
|
|
1858 |
apply(drule test)
|
|
1859 |
apply(simp)
|
|
1860 |
apply(subst (2) swap_fresh_fresh)
|
|
1861 |
apply(simp)
|
|
1862 |
apply(simp)
|
|
1863 |
apply(subst (2) swap_fresh_fresh)
|
|
1864 |
apply(simp)
|
|
1865 |
apply(simp)
|
|
1866 |
apply(subst (3) swap_fresh_fresh)
|
|
1867 |
apply(simp)
|
|
1868 |
apply(simp)
|
|
1869 |
apply(subst (3) swap_fresh_fresh)
|
|
1870 |
apply(simp)
|
|
1871 |
apply(simp)
|
|
1872 |
apply(simp)
|
|
1873 |
apply(simp add: finite_supp)
|
|
1874 |
done
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1875 |
|
2653
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1876 |
(* this should not work *)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1877 |
nominal_primrec
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1878 |
bnds :: "lam \<Rightarrow> name set"
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1879 |
where
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1880 |
"bnds (Var x) = {}"
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1881 |
| "bnds (App t1 t2) = (bnds t1) \<union> (bnds t2)"
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1882 |
| "bnds (Lam x t) = (bnds t) \<union> {x}"
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1883 |
apply(rule_tac y="x" in lam.exhaust)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1884 |
apply(simp_all)[3]
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1885 |
apply(simp_all only: lam.distinct)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1886 |
apply(simp add: lam.eq_iff)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1887 |
apply(simp add: lam.eq_iff)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1888 |
apply(simp add: lam.eq_iff)
|
d0f774d06e6e
added eqvt_at premises in function definition - however not proved at the moment
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
1889 |
sorry
|
2649
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1890 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1891 |
end
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1892 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1893 |
|
a8ebcb368a15
a modified function package where, as a test, True has been injected into the compatibility condictions
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1894 |
|