author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 12 Jan 2010 16:44:33 +0100 | |
changeset 852 | 67e5da3a356a |
parent 806 | 43336511993f |
child 870 | 2a19e0a37131 |
permissions | -rw-r--r-- |
226 | 1 |
(*notation ( output) "prop" ("#_" [1000] 1000) *) |
2 |
notation ( output) "Trueprop" ("#_" [1000] 1000) |
|
101
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
|
806
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
4 |
lemma regularize_to_injection: |
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
5 |
shows "(QUOT_TRUE l \<Longrightarrow> y) \<Longrightarrow> (l = r) \<longrightarrow> y" |
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
6 |
by(auto simp add: QUOT_TRUE_def) |
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
7 |
|
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
8 |
|
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
9 |
|
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
10 |
|
101
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
fun dest_cbinop t = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
val (t2, rhs) = Thm.dest_comb t; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
15 |
val (bop, lhs) = Thm.dest_comb t2; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
(bop, (lhs, rhs)) |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
22 |
fun dest_ceq t = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
val (bop, pair) = dest_cbinop t; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
val (bop_s, _) = Term.dest_Const (Thm.term_of bop); |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
if bop_s = "op =" then pair else (raise CTERM ("Not an equality", [t])) |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
28 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
30 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
32 |
fun split_binop_conv t = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
val (lhs, rhs) = dest_ceq t; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
val (bop, _) = dest_cbinop lhs; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
val [clT, cr2] = bop |> Thm.ctyp_of_term |> Thm.dest_ctyp; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
val [cmT, crT] = Thm.dest_ctyp cr2; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
Drule.instantiate' [SOME clT, SOME cmT, SOME crT] [NONE, NONE, NONE, NONE, SOME bop] @{thm arg_cong2} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
43 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
44 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
45 |
fun split_arg_conv t = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
46 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
47 |
val (lhs, rhs) = dest_ceq t; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
val (lop, larg) = Thm.dest_comb lhs; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
49 |
val [caT, crT] = lop |> Thm.ctyp_of_term |> Thm.dest_ctyp; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
50 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
51 |
Drule.instantiate' [SOME caT, SOME crT] [NONE, NONE, SOME lop] @{thm arg_cong} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
52 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
53 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
54 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
55 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
56 |
fun split_binop_tac n thm = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
57 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
58 |
val concl = Thm.cprem_of thm n; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
59 |
val (_, cconcl) = Thm.dest_comb concl; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
val rewr = split_binop_conv cconcl; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
61 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
62 |
rtac rewr n thm |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
63 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
64 |
handle CTERM _ => Seq.empty |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
65 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
66 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
67 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
68 |
ML {* |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
69 |
fun split_arg_tac n thm = |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
70 |
let |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
71 |
val concl = Thm.cprem_of thm n; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
72 |
val (_, cconcl) = Thm.dest_comb concl; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
73 |
val rewr = split_arg_conv cconcl; |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
74 |
in |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
75 |
rtac rewr n thm |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
end |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
77 |
handle CTERM _ => Seq.empty |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
*} |
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
80 |
|
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
81 |
lemma trueprop_cong: |
303
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
82 |
shows "(a \<equiv> b) \<Longrightarrow> (Trueprop a \<equiv> Trueprop b)" |
101
4f93c5a026d2
Reordering the code, part 3
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
83 |
by auto |
226 | 84 |
|
85 |
lemma list_induct_hol4: |
|
303
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
86 |
fixes P :: "'a list \<Rightarrow> bool" |
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
87 |
assumes a: "((P []) \<and> (\<forall>t. (P t) \<longrightarrow> (\<forall>h. (P (h # t)))))" |
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
88 |
shows "\<forall>l. (P l)" |
226 | 89 |
using a |
90 |
apply (rule_tac allI) |
|
91 |
apply (induct_tac "l") |
|
92 |
apply (simp) |
|
93 |
apply (metis) |
|
94 |
done |
|
95 |
||
301
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
96 |
ML {* |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
97 |
val no_vars = Thm.rule_attribute (fn context => fn th => |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
98 |
let |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
99 |
val ctxt = Variable.set_body false (Context.proof_of context); |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
100 |
val ((_, [th']), _) = Variable.import true [th] ctxt; |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
101 |
in th' end); |
40bb0c4718a6
Cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
226
diff
changeset
|
102 |
*} |
303
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
103 |
|
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
104 |
(*lemma equality_twice: |
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
105 |
"a = c \<Longrightarrow> b = d \<Longrightarrow> (a = b \<longrightarrow> c = d)" |
991b0e53f9dc
More code cleaning and commenting
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
301
diff
changeset
|
106 |
by auto*) |
685
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
107 |
|
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
108 |
|
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
109 |
(*interpretation code *) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
110 |
(*val bindd = ((Binding.make ("", Position.none)), ([]: Attrib.src list)) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
111 |
val ((_, [eqn1pre]), lthy5) = Variable.import true [ABS_def] lthy4; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
112 |
val eqn1i = Thm.prop_of (symmetric eqn1pre) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
113 |
val ((_, [eqn2pre]), lthy6) = Variable.import true [REP_def] lthy5; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
114 |
val eqn2i = Thm.prop_of (symmetric eqn2pre) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
115 |
|
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
116 |
val exp_morphism = ProofContext.export_morphism lthy6 (ProofContext.init (ProofContext.theory_of lthy6)); |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
117 |
val exp_term = Morphism.term exp_morphism; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
118 |
val exp = Morphism.thm exp_morphism; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
119 |
|
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
120 |
val mthd = Method.SIMPLE_METHOD ((rtac quot_thm 1) THEN |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
121 |
ALLGOALS (simp_tac (HOL_basic_ss addsimps [(symmetric (exp ABS_def)), (symmetric (exp REP_def))]))) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
122 |
val mthdt = Method.Basic (fn _ => mthd) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
123 |
val bymt = Proof.global_terminal_proof (mthdt, NONE) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
124 |
val exp_i = [(@{const_name QUOT_TYPE}, ((("QUOT_TYPE_I_" ^ (Binding.name_of qty_name)), true), |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
125 |
Expression.Named [("R", rel), ("Abs", abs), ("Rep", rep) ]))]*) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
126 |
|
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
127 |
(*||> Local_Theory.theory (fn thy => |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
128 |
let |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
129 |
val global_eqns = map exp_term [eqn2i, eqn1i]; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
130 |
(* Not sure if the following context should not be used *) |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
131 |
val (global_eqns2, lthy7) = Variable.import_terms true global_eqns lthy6; |
b12f0321dfb0
moved the interpretation code into Unused.thy
Christian Urban <urbanc@in.tum.de>
parents:
303
diff
changeset
|
132 |
val global_eqns3 = map (fn t => (bindd, t)) global_eqns2; |
806
43336511993f
Readded 'regularize_to_injection' which I believe will be needed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
685
diff
changeset
|
133 |
in ProofContext.theory_of (bymt (Expression.interpretation (exp_i, []) global_eqns3 thy)) end)*) |