| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 02 Nov 2009 15:38:49 +0100 | |
| changeset 261 | 34fb63221536 | 
| parent 259 | 22c199522bef | 
| child 263 | a159ba20979e | 
| permissions | -rw-r--r-- | 
| 0 | 1 | theory QuotMain | 
| 6 | 2 | imports QuotScript QuotList Prove | 
| 71 
35be65791f1d
exported parts of QuotMain into a separate ML-file
 Christian Urban <urbanc@in.tum.de> parents: 
70diff
changeset | 3 | uses ("quotient.ML")
 | 
| 0 | 4 | begin | 
| 5 | ||
| 261 
34fb63221536
changed Type.typ_match to Sign.typ_match
 Christian Urban <urbanc@in.tum.de> parents: 
259diff
changeset | 6 | ML {* Attrib.empty_binding *}
 | 
| 
34fb63221536
changed Type.typ_match to Sign.typ_match
 Christian Urban <urbanc@in.tum.de> parents: 
259diff
changeset | 7 | |
| 0 | 8 | locale QUOT_TYPE = | 
| 9 | fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | |
| 10 |   and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
 | |
| 11 |   and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
 | |
| 12 | assumes equiv: "EQUIV R" | |
| 13 | and rep_prop: "\<And>y. \<exists>x. Rep y = R x" | |
| 14 | and rep_inverse: "\<And>x. Abs (Rep x) = x" | |
| 15 | and abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)" | |
| 16 | and rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)" | |
| 15 | 17 | begin | 
| 0 | 18 | |
| 19 | definition | |
| 200 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 20 | ABS::"'a \<Rightarrow> 'b" | 
| 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 21 | where | 
| 0 | 22 | "ABS x \<equiv> Abs (R x)" | 
| 23 | ||
| 24 | definition | |
| 200 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 25 | REP::"'b \<Rightarrow> 'a" | 
| 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 26 | where | 
| 0 | 27 | "REP a = Eps (Rep a)" | 
| 28 | ||
| 15 | 29 | lemma lem9: | 
| 0 | 30 | shows "R (Eps (R x)) = R x" | 
| 31 | proof - | |
| 32 | have a: "R x x" using equiv by (simp add: EQUIV_REFL_SYM_TRANS REFL_def) | |
| 33 | then have "R x (Eps (R x))" by (rule someI) | |
| 15 | 34 | then show "R (Eps (R x)) = R x" | 
| 0 | 35 | using equiv unfolding EQUIV_def by simp | 
| 36 | qed | |
| 37 | ||
| 38 | theorem thm10: | |
| 24 
6885fa184e89
Merged with my changes from the morning:
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
23diff
changeset | 39 | shows "ABS (REP a) \<equiv> a" | 
| 
6885fa184e89
Merged with my changes from the morning:
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
23diff
changeset | 40 | apply (rule eq_reflection) | 
| 
6885fa184e89
Merged with my changes from the morning:
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
23diff
changeset | 41 | unfolding ABS_def REP_def | 
| 0 | 42 | proof - | 
| 15 | 43 | from rep_prop | 
| 0 | 44 | obtain x where eq: "Rep a = R x" by auto | 
| 45 | have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp | |
| 46 | also have "\<dots> = Abs (R x)" using lem9 by simp | |
| 47 | also have "\<dots> = Abs (Rep a)" using eq by simp | |
| 48 | also have "\<dots> = a" using rep_inverse by simp | |
| 49 | finally | |
| 50 | show "Abs (R (Eps (Rep a))) = a" by simp | |
| 51 | qed | |
| 52 | ||
| 15 | 53 | lemma REP_refl: | 
| 0 | 54 | shows "R (REP a) (REP a)" | 
| 55 | unfolding REP_def | |
| 56 | by (simp add: equiv[simplified EQUIV_def]) | |
| 57 | ||
| 58 | lemma lem7: | |
| 22 | 59 | shows "(R x = R y) = (Abs (R x) = Abs (R y))" | 
| 0 | 60 | apply(rule iffI) | 
| 61 | apply(simp) | |
| 62 | apply(drule rep_inject[THEN iffD2]) | |
| 63 | apply(simp add: abs_inverse) | |
| 64 | done | |
| 15 | 65 | |
| 0 | 66 | theorem thm11: | 
| 67 | shows "R r r' = (ABS r = ABS r')" | |
| 68 | unfolding ABS_def | |
| 69 | by (simp only: equiv[simplified EQUIV_def] lem7) | |
| 70 | ||
| 4 | 71 | |
| 2 | 72 | lemma REP_ABS_rsp: | 
| 4 | 73 | shows "R f (REP (ABS g)) = R f g" | 
| 74 | and "R (REP (ABS g)) f = R g f" | |
| 23 | 75 | by (simp_all add: thm10 thm11) | 
| 4 | 76 | |
| 0 | 77 | lemma QUOTIENT: | 
| 78 | "QUOTIENT R ABS REP" | |
| 79 | apply(unfold QUOTIENT_def) | |
| 80 | apply(simp add: thm10) | |
| 81 | apply(simp add: REP_refl) | |
| 82 | apply(subst thm11[symmetric]) | |
| 83 | apply(simp add: equiv[simplified EQUIV_def]) | |
| 84 | done | |
| 85 | ||
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 86 | lemma R_trans: | 
| 49 | 87 | assumes ab: "R a b" | 
| 88 | and bc: "R b c" | |
| 22 | 89 | shows "R a c" | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 90 | proof - | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 91 | have tr: "TRANS R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 92 | moreover have ab: "R a b" by fact | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 93 | moreover have bc: "R b c" by fact | 
| 22 | 94 | ultimately show "R a c" unfolding TRANS_def by blast | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 95 | qed | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 96 | |
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 97 | lemma R_sym: | 
| 49 | 98 | assumes ab: "R a b" | 
| 22 | 99 | shows "R b a" | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 100 | proof - | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 101 | have re: "SYM R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp | 
| 22 | 102 | then show "R b a" using ab unfolding SYM_def by blast | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 103 | qed | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 104 | |
| 49 | 105 | lemma R_trans2: | 
| 106 | assumes ac: "R a c" | |
| 22 | 107 | and bd: "R b d" | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 108 | shows "R a b = R c d" | 
| 200 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 109 | using ac bd | 
| 
d6a24dad5882
made quotients compatiple with Nominal; updated keyword file
 Christian Urban <urbanc@in.tum.de> parents: 
198diff
changeset | 110 | by (blast intro: R_trans R_sym) | 
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 111 | |
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 112 | lemma REPS_same: | 
| 25 
9020ee23a020
The tactic with REPEAT, CHANGED and a proper simpset.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
24diff
changeset | 113 | shows "R (REP a) (REP b) \<equiv> (a = b)" | 
| 38 | 114 | proof - | 
| 115 | have "R (REP a) (REP b) = (a = b)" | |
| 116 | proof | |
| 117 | assume as: "R (REP a) (REP b)" | |
| 118 | from rep_prop | |
| 119 | obtain x y | |
| 120 | where eqs: "Rep a = R x" "Rep b = R y" by blast | |
| 121 | from eqs have "R (Eps (R x)) (Eps (R y))" using as unfolding REP_def by simp | |
| 122 | then have "R x (Eps (R y))" using lem9 by simp | |
| 123 | then have "R (Eps (R y)) x" using R_sym by blast | |
| 124 | then have "R y x" using lem9 by simp | |
| 125 | then have "R x y" using R_sym by blast | |
| 126 | then have "ABS x = ABS y" using thm11 by simp | |
| 127 | then have "Abs (Rep a) = Abs (Rep b)" using eqs unfolding ABS_def by simp | |
| 128 | then show "a = b" using rep_inverse by simp | |
| 129 | next | |
| 130 | assume ab: "a = b" | |
| 131 | have "REFL R" using equiv EQUIV_REFL_SYM_TRANS[of R] by simp | |
| 132 | then show "R (REP a) (REP b)" unfolding REFL_def using ab by auto | |
| 133 | qed | |
| 134 | then show "R (REP a) (REP b) \<equiv> (a = b)" by simp | |
| 21 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 135 | qed | 
| 
d15121412caa
Added more useful quotient facts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
20diff
changeset | 136 | |
| 0 | 137 | end | 
| 138 | ||
| 127 
b054cf6bd179
the command "quotient" can now define more than one quotient at the same time; quotients need to be separated by and
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 139 | |
| 0 | 140 | section {* type definition for the quotient type *}
 | 
| 141 | ||
| 71 
35be65791f1d
exported parts of QuotMain into a separate ML-file
 Christian Urban <urbanc@in.tum.de> parents: 
70diff
changeset | 142 | use "quotient.ML" | 
| 0 | 143 | |
| 185 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 144 | declare [[map list = (map, LIST_REL)]] | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 145 | declare [[map * = (prod_fun, prod_rel)]] | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 146 | declare [[map "fun" = (fun_map, FUN_REL)]] | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 147 | |
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 148 | ML {* maps_lookup @{theory} "List.list" *}
 | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 149 | ML {* maps_lookup @{theory} "*" *}
 | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 150 | ML {* maps_lookup @{theory} "fun" *}
 | 
| 174 
09048a951dca
moved the map_funs setup into QuotMain
 Christian Urban <urbanc@in.tum.de> parents: 
170diff
changeset | 151 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 152 | text {* FIXME: auxiliary function *}
 | 
| 193 | 153 | ML {*
 | 
| 14 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 154 | val no_vars = Thm.rule_attribute (fn context => fn th => | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 155 | let | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 156 | val ctxt = Variable.set_body false (Context.proof_of context); | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 157 | val ((_, [th']), _) = Variable.import true [th] ctxt; | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 158 | in th' end); | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 159 | *} | 
| 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 160 | |
| 0 | 161 | section {* lifting of constants *}
 | 
| 162 | ||
| 163 | ML {*
 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 164 | fun lookup_snd _ [] _ = NONE | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 165 | | lookup_snd eq ((value, key)::xs) key' = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 166 | if eq (key', key) then SOME value | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 167 | else lookup_snd eq xs key'; | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 168 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 169 | fun lookup_qenv qenv qty = | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 170 | (case (AList.lookup (op =) qenv qty) of | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 171 | SOME rty => SOME (qty, rty) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 172 | | NONE => NONE) | 
| 185 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 173 | *} | 
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 174 | |
| 
929bc55efff7
added code for declaring map-functions
 Christian Urban <urbanc@in.tum.de> parents: 
182diff
changeset | 175 | ML {*
 | 
| 118 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 176 | (* calculates the aggregate abs and rep functions for a given type; | 
| 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 177 | repF is for constants' arguments; absF is for constants; | 
| 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 178 | function types need to be treated specially, since repF and absF | 
| 163 
3da18bf6886c
Split Finite Set example into separate file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
162diff
changeset | 179 | change | 
| 118 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 180 | *) | 
| 46 | 181 | datatype flag = absF | repF | 
| 0 | 182 | |
| 109 
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 183 | fun negF absF = repF | 
| 
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 184 | | negF repF = absF | 
| 
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 185 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 186 | fun get_fun flag qenv lthy ty = | 
| 0 | 187 | let | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 188 | |
| 0 | 189 | fun get_fun_aux s fs_tys = | 
| 190 | let | |
| 191 | val (fs, tys) = split_list fs_tys | |
| 15 | 192 | val (otys, ntys) = split_list tys | 
| 0 | 193 | val oty = Type (s, otys) | 
| 194 | val nty = Type (s, ntys) | |
| 195 | val ftys = map (op -->) tys | |
| 196 | in | |
| 130 
8e8ba210f0f7
moved the map-info and fun-info section to quotient.ML
 Christian Urban <urbanc@in.tum.de> parents: 
129diff
changeset | 197 | (case (maps_lookup (ProofContext.theory_of lthy) s) of | 
| 110 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 198 | SOME info => (list_comb (Const (#mapfun info, ftys ---> (oty --> nty)), fs), (oty, nty)) | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 199 |     | NONE      => raise ERROR ("no map association for type " ^ s))
 | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 200 | end | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 201 | |
| 118 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 202 | fun get_fun_fun fs_tys = | 
| 110 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 203 | let | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 204 | val (fs, tys) = split_list fs_tys | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 205 | val ([oty1, oty2], [nty1, nty2]) = split_list tys | 
| 128 | 206 | val oty = nty1 --> oty2 | 
| 207 | val nty = oty1 --> nty2 | |
| 110 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 208 | val ftys = map (op -->) tys | 
| 
f5f641c05794
A fix for one fun_map; doesn't work for more.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
109diff
changeset | 209 | in | 
| 118 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 210 |     (list_comb (Const (@{const_name "fun_map"}, ftys ---> oty --> nty), fs), (oty, nty))
 | 
| 0 | 211 | end | 
| 212 | ||
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 213 | fun get_const flag (qty, rty) = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 214 | let | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 215 | val thy = ProofContext.theory_of lthy | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 216 | val qty_name = Long_Name.base_name (fst (dest_Type qty)) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 217 | in | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 218 | case flag of | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 219 |       absF => (Const (Sign.full_bname thy ("ABS_" ^ qty_name), rty --> qty), (rty, qty))
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 220 |     | repF => (Const (Sign.full_bname thy ("REP_" ^ qty_name), qty --> rty), (qty, rty))
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 221 | end | 
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 222 | |
| 70 
f3cbda066c3a
consistent usage of rty (for the raw, unquotient type); tuned a bit the Isar
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 223 |   fun mk_identity ty = Abs ("", ty, Bound 0)
 | 
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 224 | |
| 0 | 225 | in | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 226 | if (AList.defined (op =) qenv ty) | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 227 | then (get_const flag (the (lookup_qenv qenv ty))) | 
| 0 | 228 | else (case ty of | 
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 229 | TFree _ => (mk_identity ty, (ty, ty)) | 
| 118 
1c30d48bfc15
slightly simplified get_fun function
 Christian Urban <urbanc@in.tum.de> parents: 
117diff
changeset | 230 | | Type (_, []) => (mk_identity ty, (ty, ty)) | 
| 109 
386671ef36bd
fixed the problem with function types; but only type_of works; cterm_of does not work
 Christian Urban <urbanc@in.tum.de> parents: 
108diff
changeset | 231 |         | Type ("fun" , [ty1, ty2]) => 
 | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 232 | get_fun_fun [get_fun (negF flag) qenv lthy ty1, get_fun flag qenv lthy ty2] | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 233 | | Type (s, tys) => get_fun_aux s (map (get_fun flag qenv lthy) tys) | 
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 234 |         | _ => raise ERROR ("no type variables")
 | 
| 0 | 235 | ) | 
| 236 | end | |
| 237 | *} | |
| 238 | ||
| 180 
fcacca9588b4
changed the definitions of liftet constants to use fun_maps
 Christian Urban <urbanc@in.tum.de> parents: 
174diff
changeset | 239 | |
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 240 | text {* produces the definition for a lifted constant *}
 | 
| 86 | 241 | |
| 2 | 242 | ML {*
 | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 243 | fun get_const_def nconst otrm qenv lthy = | 
| 0 | 244 | let | 
| 245 | val ty = fastype_of nconst | |
| 246 | val (arg_tys, res_ty) = strip_type ty | |
| 14 
5f6ee943c697
Generalized interpretation, works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
13diff
changeset | 247 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 248 | val rep_fns = map (fst o get_fun repF qenv lthy) arg_tys | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 249 | val abs_fn = (fst o get_fun absF qenv lthy) res_ty | 
| 0 | 250 | |
| 180 
fcacca9588b4
changed the definitions of liftet constants to use fun_maps
 Christian Urban <urbanc@in.tum.de> parents: 
174diff
changeset | 251 |   fun mk_fun_map (t1,t2) = Const (@{const_name "fun_map"}, dummyT) $ t1 $ t2
 | 
| 
fcacca9588b4
changed the definitions of liftet constants to use fun_maps
 Christian Urban <urbanc@in.tum.de> parents: 
174diff
changeset | 252 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 253 | val fns = Library.foldr mk_fun_map (rep_fns, abs_fn) | 
| 180 
fcacca9588b4
changed the definitions of liftet constants to use fun_maps
 Christian Urban <urbanc@in.tum.de> parents: 
174diff
changeset | 254 | |> Syntax.check_term lthy | 
| 0 | 255 | in | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 256 | fns $ otrm | 
| 0 | 257 | end | 
| 258 | *} | |
| 259 | ||
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 260 | |
| 0 | 261 | ML {*
 | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 262 | fun exchange_ty qenv ty = | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 263 | case (lookup_snd (op =) qenv ty) of | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 264 | SOME qty => qty | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 265 | | NONE => | 
| 0 | 266 | (case ty of | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 267 | Type (s, tys) => Type (s, map (exchange_ty qenv) tys) | 
| 41 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 268 | | _ => ty | 
| 
72d63aa8af68
added ctxt as explicit argument to build_goal; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
40diff
changeset | 269 | ) | 
| 0 | 270 | *} | 
| 271 | ||
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 272 | ML {* 
 | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 273 | fun ppair lthy (ty1, ty2) = | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 274 |   "(" ^ (Syntax.string_of_typ lthy ty1) ^ "," ^ (Syntax.string_of_typ lthy ty2) ^ ")"
 | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 275 | *} | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 276 | |
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 277 | ML {*
 | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 278 | fun print_env qenv lthy = | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 279 | map (ppair lthy) qenv | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 280 | |> commas | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 281 | |> tracing | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 282 | *} | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 283 | |
| 0 | 284 | ML {*
 | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 285 | fun make_const_def nconst_bname otrm mx qenv lthy = | 
| 0 | 286 | let | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 287 | val otrm_ty = fastype_of otrm | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 288 | val nconst_ty = exchange_ty qenv otrm_ty | 
| 17 
55b646c6c4cd
More naming/binding suggestions from Makarius
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
16diff
changeset | 289 | val nconst = Const (Binding.name_of nconst_bname, nconst_ty) | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 290 | val def_trm = get_const_def nconst otrm qenv lthy | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 291 | |
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 292 | val _ = print_env qenv lthy | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 293 | val _ = Syntax.string_of_typ lthy otrm_ty |> tracing | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 294 | val _ = Syntax.string_of_typ lthy nconst_ty |> tracing | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 295 | val _ = Syntax.string_of_term lthy def_trm |> tracing | 
| 0 | 296 | in | 
| 79 
c0c41fefeb06
added quotient command (you need to update isar-keywords-prove.el)
 Christian Urban <urbanc@in.tum.de> parents: 
77diff
changeset | 297 | define (nconst_bname, mx, def_trm) lthy | 
| 15 | 298 | end | 
| 0 | 299 | *} | 
| 300 | ||
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 301 | ML {*
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 302 | fun build_qenv lthy qtys = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 303 | let | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 304 |   val qenv = map (fn {qtyp, rtyp, ...} => (qtyp, rtyp)) (quotdata_lookup lthy)
 | 
| 261 
34fb63221536
changed Type.typ_match to Sign.typ_match
 Christian Urban <urbanc@in.tum.de> parents: 
259diff
changeset | 305 | val thy = ProofContext.theory_of lthy | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 306 | |
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 307 | fun find_assoc ((qty', rty')::rest) qty = | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 308 | let | 
| 261 
34fb63221536
changed Type.typ_match to Sign.typ_match
 Christian Urban <urbanc@in.tum.de> parents: 
259diff
changeset | 309 | val inst = Sign.typ_match thy (qty', qty) Vartab.empty | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 310 | in | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 311 | (Envir.norm_type inst qty, Envir.norm_type inst rty') | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 312 | end | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 313 | | find_assoc [] qty = | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 314 | let | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 315 | val qtystr = quote (Syntax.string_of_typ lthy qty) | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 316 | in | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 317 | error (implode ["Quotient type ", qtystr, " does not exists"]) | 
| 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 318 | end | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 319 | in | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 320 | map (find_assoc qenv) qtys | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 321 | end | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 322 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 323 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 324 | ML {*
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 325 | val qd_parser = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 326 | (Args.parens (OuterParse.$$$ "for" |-- (Scan.repeat OuterParse.typ))) -- | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 327 | (SpecParse.constdecl -- (SpecParse.opt_thm_name ":" -- OuterParse.prop)) | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 328 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 329 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 330 | ML {*
 | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 331 | fun parse_qd_spec (qtystrs, ((bind, tystr, mx), (attr__, propstr))) lthy = | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 332 | let | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 333 | val qtys = map (Syntax.check_typ lthy o Syntax.parse_typ lthy) qtystrs | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 334 | val qenv = build_qenv lthy qtys | 
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 335 | val qty = Syntax.parse_typ lthy (the tystr) |> Syntax.check_typ lthy | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 336 | val prop = Syntax.parse_prop lthy propstr |> Syntax.check_prop lthy | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 337 | val (lhs, rhs) = Logic.dest_equals prop | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 338 | in | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 339 | make_const_def bind rhs mx qenv lthy |> snd | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 340 | end | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 341 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 342 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 343 | ML {*
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 344 | val _ = OuterSyntax.local_theory "quotient_def" "lifted definition of constants" | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 345 | OuterKeyword.thy_decl (qd_parser >> parse_qd_spec) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 346 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 347 | |
| 254 
77ff9624cfd6
fixed the problem with types in map
 Christian Urban <urbanc@in.tum.de> parents: 
252diff
changeset | 348 | |
| 139 | 349 | section {* ATOMIZE *}
 | 
| 350 | ||
| 351 | text {*
 | |
| 352 | Unabs_def converts a definition given as | |
| 353 | ||
| 354 | c \<equiv> %x. %y. f x y | |
| 355 | ||
| 356 | to a theorem of the form | |
| 357 | ||
| 358 | c x y \<equiv> f x y | |
| 359 | ||
| 360 | This function is needed to rewrite the right-hand | |
| 361 | side to the left-hand side. | |
| 362 | *} | |
| 363 | ||
| 364 | ML {*
 | |
| 365 | fun unabs_def ctxt def = | |
| 366 | let | |
| 367 | val (lhs, rhs) = Thm.dest_equals (cprop_of def) | |
| 368 | val xs = strip_abs_vars (term_of rhs) | |
| 369 | val (_, ctxt') = Variable.add_fixes (map fst xs) ctxt | |
| 370 | ||
| 371 | val thy = ProofContext.theory_of ctxt' | |
| 372 | val cxs = map (cterm_of thy o Free) xs | |
| 373 | val new_lhs = Drule.list_comb (lhs, cxs) | |
| 374 | ||
| 375 | fun get_conv [] = Conv.rewr_conv def | |
| 376 | | get_conv (x::xs) = Conv.fun_conv (get_conv xs) | |
| 377 | in | |
| 378 | get_conv xs new_lhs |> | |
| 379 | singleton (ProofContext.export ctxt' ctxt) | |
| 380 | end | |
| 381 | *} | |
| 382 | ||
| 383 | lemma atomize_eqv[atomize]: | |
| 384 | shows "(Trueprop A \<equiv> Trueprop B) \<equiv> (A \<equiv> B)" | |
| 385 | proof | |
| 386 | assume "A \<equiv> B" | |
| 387 | then show "Trueprop A \<equiv> Trueprop B" by unfold | |
| 388 | next | |
| 389 | assume *: "Trueprop A \<equiv> Trueprop B" | |
| 390 | have "A = B" | |
| 391 | proof (cases A) | |
| 392 | case True | |
| 393 | have "A" by fact | |
| 394 | then show "A = B" using * by simp | |
| 395 | next | |
| 396 | case False | |
| 397 | have "\<not>A" by fact | |
| 398 | then show "A = B" using * by auto | |
| 399 | qed | |
| 400 | then show "A \<equiv> B" by (rule eq_reflection) | |
| 401 | qed | |
| 402 | ||
| 403 | ML {*
 | |
| 404 | fun atomize_thm thm = | |
| 405 | let | |
| 221 | 406 | val thm' = Thm.freezeT (forall_intr_vars thm) | 
| 139 | 407 | val thm'' = ObjectLogic.atomize (cprop_of thm') | 
| 408 | in | |
| 221 | 409 |   @{thm Pure.equal_elim_rule1} OF [thm'', thm']
 | 
| 139 | 410 | end | 
| 411 | *} | |
| 412 | ||
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 413 | ML {* atomize_thm @{thm list.induct} *}
 | 
| 139 | 414 | |
| 415 | section {* REGULARIZE *}
 | |
| 416 | ||
| 417 | text {* tyRel takes a type and builds a relation that a quantifier over this
 | |
| 418 | type needs to respect. *} | |
| 419 | ML {*
 | |
| 257 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 420 | fun matches (ty1, ty2) = | 
| 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 421 | Type.raw_instance (ty1, Logic.varifyT ty2); | 
| 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 422 | |
| 139 | 423 | fun tyRel ty rty rel lthy = | 
| 257 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 424 | if matches (rty, ty) | 
| 139 | 425 | then rel | 
| 426 | else (case ty of | |
| 427 | Type (s, tys) => | |
| 428 | let | |
| 429 |               val tys_rel = map (fn ty => ty --> ty --> @{typ bool}) tys;
 | |
| 430 |               val ty_out = ty --> ty --> @{typ bool};
 | |
| 431 | val tys_out = tys_rel ---> ty_out; | |
| 432 | in | |
| 433 | (case (maps_lookup (ProofContext.theory_of lthy) s) of | |
| 434 | SOME (info) => list_comb (Const (#relfun info, tys_out), map (fn ty => tyRel ty rty rel lthy) tys) | |
| 435 | | NONE => HOLogic.eq_const ty | |
| 436 | ) | |
| 437 | end | |
| 438 | | _ => HOLogic.eq_const ty) | |
| 439 | *} | |
| 440 | ||
| 441 | definition | |
| 442 |   Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 | |
| 443 | where | |
| 444 | "(x \<in> p) \<Longrightarrow> (Babs p m x = m x)" | |
| 445 | (* TODO: Consider defining it with an "if"; sth like: | |
| 446 | Babs p m = \<lambda>x. if x \<in> p then m x else undefined | |
| 447 | *) | |
| 448 | ||
| 449 | ML {*
 | |
| 450 | fun needs_lift (rty as Type (rty_s, _)) ty = | |
| 451 | case ty of | |
| 452 | Type (s, tys) => | |
| 453 | (s = rty_s) orelse (exists (needs_lift rty) tys) | |
| 454 | | _ => false | |
| 455 | ||
| 456 | *} | |
| 457 | ||
| 458 | ML {*
 | |
| 459 | fun mk_babs ty ty' = Const (@{const_name "Babs"}, [ty' --> @{typ bool}, ty] ---> ty)
 | |
| 460 | fun mk_ball ty = Const (@{const_name "Ball"}, [ty, ty] ---> @{typ bool})
 | |
| 461 | fun mk_bex ty = Const (@{const_name "Bex"}, [ty, ty] ---> @{typ bool})
 | |
| 462 | fun mk_resp ty = Const (@{const_name Respects}, [[ty, ty] ---> @{typ bool}, ty] ---> @{typ bool})
 | |
| 463 | *} | |
| 464 | ||
| 465 | (* applies f to the subterm of an abstractions, otherwise to the given term *) | |
| 466 | ML {*
 | |
| 467 | fun apply_subt f trm = | |
| 468 | case trm of | |
| 145 | 469 | Abs (x, T, t) => | 
| 470 | let | |
| 471 | val (x', t') = Term.dest_abs (x, T, t) | |
| 472 | in | |
| 473 | Term.absfree (x', T, f t') | |
| 474 | end | |
| 139 | 475 | | _ => f trm | 
| 476 | *} | |
| 477 | ||
| 478 | (* FIXME: if there are more than one quotient, then you have to look up the relation *) | |
| 479 | ML {*
 | |
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 480 | fun my_reg lthy rel rty trm = | 
| 139 | 481 | case trm of | 
| 482 | Abs (x, T, t) => | |
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 483 | if (needs_lift rty T) then | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 484 | let | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 485 | val rrel = tyRel T rty rel lthy | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 486 | in | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 487 | (mk_babs (fastype_of trm) T) $ (mk_resp T $ rrel) $ (apply_subt (my_reg lthy rel rty) trm) | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 488 | end | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 489 | else | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 490 | Abs(x, T, (apply_subt (my_reg lthy rel rty) t)) | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 491 |   | Const (@{const_name "All"}, ty) $ (t as Abs (x, T, _)) =>
 | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 492 | let | 
| 139 | 493 | val ty1 = domain_type ty | 
| 494 | val ty2 = domain_type ty1 | |
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 495 | val rrel = tyRel T rty rel lthy | 
| 139 | 496 | in | 
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 497 | if (needs_lift rty T) then | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 498 | (mk_ball ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t) | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 499 | else | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 500 |            Const (@{const_name "All"}, ty) $ apply_subt (my_reg lthy rel rty) t
 | 
| 139 | 501 | end | 
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 502 |   | Const (@{const_name "Ex"}, ty) $ (t as Abs (x, T, _)) =>
 | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 503 | let | 
| 139 | 504 | val ty1 = domain_type ty | 
| 505 | val ty2 = domain_type ty1 | |
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 506 | val rrel = tyRel T rty rel lthy | 
| 139 | 507 | in | 
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 508 | if (needs_lift rty T) then | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 509 | (mk_bex ty1) $ (mk_resp ty2 $ rrel) $ (apply_subt (my_reg lthy rel rty) t) | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 510 | else | 
| 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 511 |            Const (@{const_name "Ex"}, ty) $ apply_subt (my_reg lthy rel rty) t
 | 
| 139 | 512 | end | 
| 251 | 513 |   | Const (@{const_name "op ="}, ty) $ t =>
 | 
| 514 | if needs_lift rty (fastype_of t) then | |
| 515 | (tyRel (fastype_of t) rty rel lthy) $ t | |
| 516 |       else Const (@{const_name "op ="}, ty) $ (my_reg lthy rel rty t)
 | |
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 517 | | t1 $ t2 => (my_reg lthy rel rty t1) $ (my_reg lthy rel rty t2) | 
| 139 | 518 | | _ => trm | 
| 519 | *} | |
| 520 | ||
| 141 
0ffc37761e53
Further reorganization
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
140diff
changeset | 521 | text {* Assumes that the given theorem is atomized *}
 | 
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 522 | ML {*
 | 
| 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 523 | fun build_regularize_goal thm rty rel lthy = | 
| 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 524 | Logic.mk_implies | 
| 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 525 | ((prop_of thm), | 
| 248 
6ed87b3d358c
Finally merged the code of the versions of regularize and tested examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
241diff
changeset | 526 | (my_reg lthy rel rty (prop_of thm))) | 
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 527 | *} | 
| 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 528 | |
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 529 | lemma universal_twice: "(\<And>x. (P x \<longrightarrow> Q x)) \<Longrightarrow> ((\<forall>x. P x) \<longrightarrow> (\<forall>x. Q x))" | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 530 | apply (auto) | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 531 | done | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 532 | |
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 533 | lemma implication_twice: "(c \<longrightarrow> a) \<Longrightarrow> (a \<Longrightarrow> b \<longrightarrow> d) \<Longrightarrow> (a \<longrightarrow> b) \<longrightarrow> (c \<longrightarrow> d)" | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 534 | apply (auto) | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 535 | done | 
| 251 | 536 | |
| 187 | 537 | ML {*
 | 
| 251 | 538 | fun regularize thm rty rel rel_eqv rel_refl lthy = | 
| 187 | 539 | let | 
| 540 | val g = build_regularize_goal thm rty rel lthy; | |
| 541 | fun tac ctxt = | |
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 542 | (ObjectLogic.full_atomize_tac) THEN' | 
| 251 | 543 | REPEAT_ALL_NEW (FIRST' [ | 
| 544 | rtac rel_refl, | |
| 545 | atac, | |
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 546 |       rtac @{thm universal_twice},
 | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 547 |       (rtac @{thm impI} THEN' atac),
 | 
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 548 |       rtac @{thm implication_twice},
 | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 549 | EqSubst.eqsubst_tac ctxt [0] | 
| 187 | 550 |         [(@{thm equiv_res_forall} OF [rel_eqv]),
 | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 551 |          (@{thm equiv_res_exists} OF [rel_eqv])],
 | 
| 252 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 552 |       (rtac @{thm impI} THEN' (asm_full_simp_tac (Simplifier.context ctxt HOL_ss)) THEN' rtac rel_refl),
 | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 553 |       (rtac @{thm RIGHT_RES_FORALL_REGULAR})
 | 
| 
e30997c88050
Regularize for equalities and a better tactic. "alpha.cases" now lifts.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
251diff
changeset | 554 | ]); | 
| 187 | 555 | val cthm = Goal.prove lthy [] [] g (fn x => tac (#context x) 1); | 
| 556 | in | |
| 557 | cthm OF [thm] | |
| 558 | end | |
| 559 | *} | |
| 560 | ||
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 561 | section {* RepAbs injection *}
 | 
| 139 | 562 | |
| 161 
2ee03759a22f
Trying to get a simpler lemma with the whole infrastructure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
160diff
changeset | 563 | (* Needed to have a meta-equality *) | 
| 
2ee03759a22f
Trying to get a simpler lemma with the whole infrastructure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
160diff
changeset | 564 | lemma id_def_sym: "(\<lambda>x. x) \<equiv> id" | 
| 
2ee03759a22f
Trying to get a simpler lemma with the whole infrastructure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
160diff
changeset | 565 | by (simp add: id_def) | 
| 
2ee03759a22f
Trying to get a simpler lemma with the whole infrastructure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
160diff
changeset | 566 | |
| 139 | 567 | ML {*
 | 
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 568 | fun old_exchange_ty rty qty ty = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 569 | if ty = rty | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 570 | then qty | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 571 | else | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 572 | (case ty of | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 573 | Type (s, tys) => Type (s, map (old_exchange_ty rty qty) tys) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 574 | | _ => ty | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 575 | ) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 576 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 577 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 578 | ML {*
 | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 579 | fun old_get_fun flag rty qty lthy ty = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 580 | get_fun flag [(qty, rty)] lthy ty | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 581 | |
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 582 | fun old_make_const_def nconst_bname otrm mx rty qty lthy = | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 583 | make_const_def nconst_bname otrm mx [(qty, rty)] lthy | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 584 | *} | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 585 | |
| 235 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 586 | text {* Does the same as 'subst' in a given prop or theorem *}
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 587 | ML {*
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 588 | fun eqsubst_prop ctxt thms t = | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 589 | let | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 590 | val goalstate = Goal.init (cterm_of (ProofContext.theory_of ctxt) t) | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 591 | val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 592 | NONE => error "eqsubst_prop" | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 593 | | SOME th => cprem_of th 1 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 594 | in term_of a' end | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 595 | *} | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 596 | |
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 597 | ML {*
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 598 | fun repeat_eqsubst_prop ctxt thms t = | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 599 | repeat_eqsubst_prop ctxt thms (eqsubst_prop ctxt thms t) | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 600 | handle _ => t | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 601 | *} | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 602 | |
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 603 | |
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 604 | ML {*
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 605 | fun eqsubst_thm ctxt thms thm = | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 606 | let | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 607 | val goalstate = Goal.init (Thm.cprop_of thm) | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 608 | val a' = case (SINGLE (EqSubst.eqsubst_tac ctxt [0] thms 1) goalstate) of | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 609 | NONE => error "eqsubst_thm" | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 610 | | SOME th => cprem_of th 1 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 611 | val tac = (EqSubst.eqsubst_tac ctxt [0] thms 1) THEN simp_tac HOL_ss 1 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 612 | val cgoal = cterm_of (ProofContext.theory_of ctxt) (Logic.mk_equals (term_of (Thm.cprop_of thm), term_of a')) | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 613 | val rt = Toplevel.program (fn () => Goal.prove_internal [] cgoal (fn _ => tac)); | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 614 | in | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 615 |     @{thm Pure.equal_elim_rule1} OF [rt,thm]
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 616 | end | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 617 | *} | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 618 | |
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 619 | ML {*
 | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 620 | fun repeat_eqsubst_thm ctxt thms thm = | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 621 | repeat_eqsubst_thm ctxt thms (eqsubst_thm ctxt thms thm) | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 622 | handle _ => thm | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 623 | *} | 
| 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 624 | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 625 | ML {*
 | 
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 626 | fun build_repabs_term lthy thm constructors rty qty = | 
| 139 | 627 | let | 
| 628 | fun mk_rep tm = | |
| 629 | let | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 630 | val ty = old_exchange_ty rty qty (fastype_of tm) | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 631 | in fst (old_get_fun repF rty qty lthy ty) $ tm end | 
| 139 | 632 | |
| 633 | fun mk_abs tm = | |
| 634 | let | |
| 218 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 635 | val ty = old_exchange_ty rty qty (fastype_of tm) in | 
| 
df05cd030d2f
added infrastructure for defining lifted constants
 Christian Urban <urbanc@in.tum.de> parents: 
214diff
changeset | 636 | fst (old_get_fun absF rty qty lthy ty) $ tm end | 
| 139 | 637 | |
| 638 | fun is_constructor (Const (x, _)) = member (op =) constructors x | |
| 639 | | is_constructor _ = false; | |
| 640 | ||
| 641 | fun build_aux lthy tm = | |
| 642 | case tm of | |
| 643 | Abs (a as (_, vty, _)) => | |
| 644 | let | |
| 645 | val (vs, t) = Term.dest_abs a; | |
| 646 | val v = Free(vs, vty); | |
| 647 | val t' = lambda v (build_aux lthy t) | |
| 648 | in | |
| 649 | if (not (needs_lift rty (fastype_of tm))) then t' | |
| 650 | else mk_rep (mk_abs ( | |
| 651 | if not (needs_lift rty vty) then t' | |
| 652 | else | |
| 653 | let | |
| 654 | val v' = mk_rep (mk_abs v); | |
| 655 | val t1 = Envir.beta_norm (t' $ v') | |
| 656 | in | |
| 657 | lambda v t1 | |
| 658 | end | |
| 659 | )) | |
| 660 | end | |
| 661 | | x => | |
| 662 | let | |
| 663 | val (opp, tms0) = Term.strip_comb tm | |
| 664 | val tms = map (build_aux lthy) tms0 | |
| 665 | val ty = fastype_of tm | |
| 666 | in | |
| 667 |         if (((fst (Term.dest_Const opp)) = @{const_name Respects}) handle _ => false)
 | |
| 668 | then (list_comb (opp, (hd tms0) :: (tl tms))) | |
| 669 | else if (is_constructor opp andalso needs_lift rty ty) then | |
| 670 | mk_rep (mk_abs (list_comb (opp,tms))) | |
| 671 | else if ((Term.is_Free opp) andalso (length tms > 0) andalso (needs_lift rty ty)) then | |
| 149 
7cf1d7adfc5f
Removed some debugging messages
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
148diff
changeset | 672 | mk_rep(mk_abs(list_comb(opp,tms))) | 
| 139 | 673 | else if tms = [] then opp | 
| 674 | else list_comb(opp, tms) | |
| 675 | end | |
| 676 | in | |
| 235 
7affee8f90f5
Using subst for identity definition.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
223diff
changeset | 677 |     repeat_eqsubst_prop lthy @{thms id_def_sym}
 | 
| 161 
2ee03759a22f
Trying to get a simpler lemma with the whole infrastructure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
160diff
changeset | 678 | (build_aux lthy (Thm.prop_of thm)) | 
| 139 | 679 | end | 
| 680 | *} | |
| 681 | ||
| 141 
0ffc37761e53
Further reorganization
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
140diff
changeset | 682 | text {* Assumes that it is given a regularized theorem *}
 | 
| 139 | 683 | ML {*
 | 
| 140 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 684 | fun build_repabs_goal ctxt thm cons rty qty = | 
| 
00d141f2daa7
Further reorganizing the file
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
139diff
changeset | 685 | Logic.mk_equals ((Thm.prop_of thm), (build_repabs_term ctxt thm cons rty qty)) | 
| 139 | 686 | *} | 
| 687 | ||
| 187 | 688 | ML {*
 | 
| 689 | fun instantiate_tac thm = Subgoal.FOCUS (fn {concl, ...} =>
 | |
| 690 | let | |
| 691 | val pat = Drule.strip_imp_concl (cprop_of thm) | |
| 692 | val insts = Thm.match (pat, concl) | |
| 693 | in | |
| 694 | rtac (Drule.instantiate insts thm) 1 | |
| 152 
53277fbb2dba
Simplified the proof with some tactic... Still hangs sometimes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
151diff
changeset | 695 | end | 
| 187 | 696 | handle _ => no_tac | 
| 697 | ) | |
| 698 | *} | |
| 699 | ||
| 700 | ML {*
 | |
| 701 | fun res_forall_rsp_tac ctxt = | |
| 702 |   (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
 | |
| 703 |   THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
 | |
| 704 |   THEN' instantiate_tac @{thm RES_FORALL_RSP} ctxt THEN'
 | |
| 705 |   (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
 | |
| 706 | *} | |
| 707 | ||
| 708 | ML {*
 | |
| 709 | fun res_exists_rsp_tac ctxt = | |
| 710 |   (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
 | |
| 711 |   THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI}
 | |
| 712 |   THEN' instantiate_tac @{thm RES_EXISTS_RSP} ctxt THEN'
 | |
| 713 |   (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
 | |
| 714 | *} | |
| 715 | ||
| 716 | ||
| 717 | ML {*
 | |
| 718 | fun quotient_tac quot_thm = | |
| 719 | REPEAT_ALL_NEW (FIRST' [ | |
| 720 |     rtac @{thm FUN_QUOTIENT},
 | |
| 721 | rtac quot_thm, | |
| 722 |     rtac @{thm IDENTITY_QUOTIENT}
 | |
| 723 | ]) | |
| 724 | *} | |
| 725 | ||
| 726 | ML {*
 | |
| 727 | fun LAMBDA_RES_TAC ctxt i st = | |
| 728 | (case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of | |
| 729 | (_ $ (_ $ (Abs(_,_,_))$(Abs(_,_,_)))) => | |
| 730 |       (EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
 | |
| 731 |       (rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
 | |
| 732 | | _ => fn _ => no_tac) i st | |
| 733 | *} | |
| 734 | ||
| 735 | ML {*
 | |
| 736 | fun WEAK_LAMBDA_RES_TAC ctxt i st = | |
| 737 | (case (term_of o #concl o fst) (Subgoal.focus ctxt i st) of | |
| 738 | (_ $ (_ $ _$(Abs(_,_,_)))) => | |
| 739 |       (EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
 | |
| 740 |       (rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
 | |
| 741 | | (_ $ (_ $ (Abs(_,_,_))$_)) => | |
| 742 |       (EqSubst.eqsubst_tac ctxt [0] @{thms FUN_REL.simps}) THEN'
 | |
| 743 |       (rtac @{thm allI}) THEN' (rtac @{thm allI}) THEN' (rtac @{thm impI})
 | |
| 744 | | _ => fn _ => no_tac) i st | |
| 745 | *} | |
| 746 | ||
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 747 | ML {*
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 748 | fun APPLY_RSP_TAC rty = Subgoal.FOCUS (fn {concl, ...} =>
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 749 | let | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 750 | val (_ $ (R $ (f $ _) $ (_ $ _))) = term_of concl; | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 751 |     val pat = Drule.strip_imp_concl (cprop_of @{thm APPLY_RSP});
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 752 | val insts = Thm.match (pat, concl) | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 753 | in | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 754 | if needs_lift rty (type_of f) then | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 755 |     rtac (Drule.instantiate insts @{thm APPLY_RSP}) 1
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 756 | else no_tac | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 757 | end | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 758 | handle _ => no_tac) | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 759 | *} | 
| 187 | 760 | |
| 761 | ML {*
 | |
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 762 | fun r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms = | 
| 187 | 763 | (FIRST' [ | 
| 259 | 764 | (*    rtac @{thm FUN_QUOTIENT},
 | 
| 187 | 765 | rtac quot_thm, | 
| 259 | 766 |     rtac @{thm IDENTITY_QUOTIENT},*)
 | 
| 187 | 767 | rtac trans_thm, | 
| 768 | LAMBDA_RES_TAC ctxt, | |
| 769 | res_forall_rsp_tac ctxt, | |
| 770 | res_exists_rsp_tac ctxt, | |
| 771 | ( | |
| 772 | (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps rsp_thms)) | |
| 773 | THEN_ALL_NEW (fn _ => no_tac) | |
| 774 | ), | |
| 775 |     (instantiate_tac @{thm REP_ABS_RSP(1)} ctxt THEN' (RANGE [quotient_tac quot_thm])),
 | |
| 776 | rtac refl, | |
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 777 | (*    rtac @{thm arg_cong2[of _ _ _ _ "op ="]},*)
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 778 | (APPLY_RSP_TAC rty ctxt THEN' (RANGE [quotient_tac quot_thm, quotient_tac quot_thm])), | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 779 |     Cong_Tac.cong_tac @{thm cong},
 | 
| 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 780 |     rtac @{thm ext},
 | 
| 187 | 781 | rtac reflex_thm, | 
| 782 | atac, | |
| 783 | ( | |
| 784 |      (simp_tac ((Simplifier.context ctxt HOL_ss) addsimps @{thms FUN_REL.simps}))
 | |
| 785 | THEN_ALL_NEW (fn _ => no_tac) | |
| 786 | ), | |
| 787 | WEAK_LAMBDA_RES_TAC ctxt | |
| 788 | ]) | |
| 789 | *} | |
| 790 | ||
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 791 | ML {*
 | 
| 210 
f88ea69331bf
Simplfied interface to repabs_injection.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
209diff
changeset | 792 | fun repabs lthy thm constructors rty qty quot_thm reflex_thm trans_thm rsp_thms = | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 793 | let | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 794 | val rt = build_repabs_term lthy thm constructors rty qty; | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 795 | val rg = Logic.mk_equals ((Thm.prop_of thm), rt); | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 796 | fun tac ctxt = (ObjectLogic.full_atomize_tac) THEN' | 
| 206 
1e227c9ee915
Fixed APPLY_RSP vs Cong in the InjRepAbs tactic.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
200diff
changeset | 797 | (REPEAT_ALL_NEW (r_mk_comb_tac ctxt rty quot_thm reflex_thm trans_thm rsp_thms)); | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 798 | val cthm = Goal.prove lthy [] [] rg (fn x => tac (#context x) 1); | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 799 | in | 
| 210 
f88ea69331bf
Simplfied interface to repabs_injection.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
209diff
changeset | 800 |     @{thm Pure.equal_elim_rule1} OF [cthm, thm]
 | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 801 | end | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 802 | *} | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 803 | |
| 187 | 804 | section {* Cleaning the goal *}
 | 
| 805 | ||
| 236 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 806 | lemma prod_fun_id: "prod_fun id id = id" | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 807 | apply (simp add: prod_fun_def) | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 808 | done | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 809 | lemma map_id: "map id x = x" | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 810 | apply (induct x) | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 811 | apply (simp_all add: map.simps) | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 812 | done | 
| 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 813 | |
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 814 | text {* expects atomized definition *}
 | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 815 | ML {*
 | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 816 | fun add_lower_defs_aux ctxt thm = | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 817 | let | 
| 209 
1e8e1b736586
map_append lifted automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
207diff
changeset | 818 |       val e1 = @{thm fun_cong} OF [thm];
 | 
| 
1e8e1b736586
map_append lifted automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
207diff
changeset | 819 |       val f = eqsubst_thm ctxt @{thms fun_map.simps} e1;
 | 
| 
1e8e1b736586
map_append lifted automatically.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
207diff
changeset | 820 |       val g = MetaSimplifier.rewrite_rule @{thms id_def_sym} f;
 | 
| 236 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 821 |       val h = repeat_eqsubst_thm ctxt @{thms FUN_MAP_I id_apply prod_fun_id map_id} g
 | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 822 | in | 
| 236 
23f9fead8bd6
Cleaning of 'map id' and 'prod_fun id id' in lower_defs.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
235diff
changeset | 823 | thm :: (add_lower_defs_aux ctxt h) | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 824 | end | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 825 | handle _ => [thm] | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 826 | *} | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 827 | |
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 828 | ML {*
 | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 829 | fun add_lower_defs ctxt defs = | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 830 | let | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 831 | val defs_pre_sym = map symmetric defs | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 832 | val defs_atom = map atomize_thm defs_pre_sym | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 833 | val defs_all = flat (map (add_lower_defs_aux ctxt) defs_atom) | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 834 | in | 
| 214 | 835 | map Thm.varifyT defs_all | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 836 | end | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 837 | *} | 
| 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 838 | |
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 839 | ML {*
 | 
| 241 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 840 | fun findabs_all rty tm = | 
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 841 | case tm of | 
| 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 842 | Abs(_, T, b) => | 
| 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 843 | let | 
| 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 844 |           val b' = subst_bound ((Free ("x", T)), b);
 | 
| 241 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 845 | val tys = findabs_all rty b' | 
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 846 | val ty = fastype_of tm | 
| 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 847 | in if needs_lift rty ty then (ty :: tys) else tys | 
| 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 848 | end | 
| 241 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 849 | | f $ a => (findabs_all rty f) @ (findabs_all rty a) | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 850 | | _ => []; | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 851 | fun findabs rty tm = distinct (op =) (findabs_all rty tm) | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 852 | *} | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 853 | |
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 854 | |
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 855 | ML {*
 | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 856 | fun findaps_all rty tm = | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 857 | case tm of | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 858 | Abs(_, T, b) => | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 859 |         findaps_all rty (subst_bound ((Free ("x", T)), b))
 | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 860 | | (f $ a) => (findaps_all rty f @ findaps_all rty a) | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 861 |     | Free (_, (T as (Type ("fun", (_ :: _))))) => (if needs_lift rty T then [T] else [])
 | 
| 241 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 862 | | _ => []; | 
| 
60acf3d3a4a0
Finding applications and duplicates filtered out in abstractions
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
239diff
changeset | 863 | fun findaps rty tm = distinct (op =) (findaps_all rty tm) | 
| 191 
b97f3f5fbc18
Symmetry of integer addition
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
190diff
changeset | 864 | *} | 
| 190 
ca1a24aa822e
Finished the code for adding lower defs, and more things moved to QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
187diff
changeset | 865 | |
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 866 | ML {*
 | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 867 | fun make_simp_prs_thm lthy quot_thm thm typ = | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 868 | let | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 869 | val (_, [lty, rty]) = dest_Type typ; | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 870 | val thy = ProofContext.theory_of lthy; | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 871 | val (lcty, rcty) = (ctyp_of thy lty, ctyp_of thy rty) | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 872 | val inst = [SOME lcty, NONE, SOME rcty]; | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 873 | val lpi = Drule.instantiate' inst [] thm; | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 874 | val tac = | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 875 | (compose_tac (false, lpi, 2)) THEN_ALL_NEW | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 876 | (quotient_tac quot_thm); | 
| 259 | 877 | val gc = Drule.strip_imp_concl (cprop_of lpi); | 
| 878 | val t = Goal.prove_internal [] gc (fn _ => tac 1) | |
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 879 | in | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 880 |     MetaSimplifier.rewrite_rule [@{thm eq_reflection} OF @{thms id_apply}] t
 | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 881 | end | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 882 | *} | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 883 | |
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 884 | ML {*
 | 
| 259 | 885 | fun simp_allex_prs quot thms thm = | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 886 | let | 
| 259 | 887 |       val ts = [@{thm FORALL_PRS} OF [quot], @{thm EXISTS_PRS} OF [quot]] @ thms
 | 
| 888 |       val sym_ts = map (fn x => @{thm "HOL.sym"} OF [x]) ts;
 | |
| 889 |       val eq_ts = map (fn x => @{thm "eq_reflection"} OF [x]) sym_ts
 | |
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 890 | in | 
| 259 | 891 | MetaSimplifier.rewrite_rule eq_ts thm | 
| 197 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 892 | end | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 893 | *} | 
| 
c0f2db9a243b
Further reordering in Int code.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
193diff
changeset | 894 | |
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 895 | ML {*
 | 
| 239 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 896 | fun lookup_quot_data lthy qty = | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 897 | let | 
| 257 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 898 | val SOME quotdata = find_first (fn x => matches ((#qtyp x), qty)) (quotdata_lookup lthy) | 
| 
68bd5c2a1b96
Fixed quotdata_lookup.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
255diff
changeset | 899 | val rty = Logic.unvarifyT (#rtyp quotdata) | 
| 239 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 900 | val rel = #rel quotdata | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 901 | val rel_eqv = #equiv_thm quotdata | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 902 |     val rel_refl_pre = @{thm EQUIV_REFL} OF [rel_eqv]
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 903 |     val rel_refl = @{thm spec} OF [MetaSimplifier.rewrite_rule [@{thm REFL_def}] rel_refl_pre]
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 904 | in | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 905 | (rty, rel, rel_refl, rel_eqv) | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 906 | end | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 907 | *} | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 908 | |
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 909 | ML {*
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 910 | fun lookup_quot_thms lthy qty_name = | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 911 | let | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 912 | val thy = ProofContext.theory_of lthy; | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 913 |     val trans2 = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".R_trans2")
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 914 |     val reps_same = PureThy.get_thm thy ("QUOT_TYPE_I_" ^ qty_name ^ ".REPS_same")
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 915 |     val quot = PureThy.get_thm thy ("QUOTIENT_" ^ qty_name)
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 916 | in | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 917 | (trans2, reps_same, quot) | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 918 | end | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 919 | *} | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 920 | |
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 921 | ML {*
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 922 | fun lookup_quot_consts defs = | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 923 | let | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 924 | fun dest_term (a $ b) = (a, b); | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 925 | val def_terms = map (snd o Logic.dest_equals o concl_of) defs; | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 926 | in | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 927 | map (fst o dest_Const o snd o dest_term) def_terms | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 928 | end | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 929 | *} | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 930 | |
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 931 | ML {*
 | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 932 | fun lift_thm lthy qty qty_name rsp_thms defs t = let | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 933 | val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data lthy qty; | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 934 | val (trans2, reps_same, quot) = lookup_quot_thms lthy qty_name; | 
| 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 935 | val consts = lookup_quot_consts defs; | 
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 936 | val t_a = atomize_thm t; | 
| 251 | 937 | val t_r = regularize t_a rty rel rel_eqv rel_refl lthy; | 
| 210 
f88ea69331bf
Simplfied interface to repabs_injection.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
209diff
changeset | 938 | val t_t = repabs lthy t_r consts rty qty quot rel_refl trans2 rsp_thms; | 
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 939 | val abs = findabs rty (prop_of t_a); | 
| 253 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 940 | val aps = findaps rty (prop_of t_a); | 
| 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 941 |   val app_prs_thms = map (make_simp_prs_thm lthy quot @{thm APP_PRS}) aps;
 | 
| 
e169a99c6ada
Automatic computation of application preservation and manually finished "alpha.induct". Slow...
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
252diff
changeset | 942 |   val lam_prs_thms = map (make_simp_prs_thm lthy quot @{thm LAMBDA_PRS}) abs;
 | 
| 259 | 943 | val t_a = simp_allex_prs quot [] t_t; | 
| 944 | val t_l = repeat_eqsubst_thm lthy (lam_prs_thms @ app_prs_thms) t_a; | |
| 239 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 945 | val defs_sym = add_lower_defs lthy defs; | 
| 259 | 946 | val defs_sym_eq = map (fn x => eq_reflection OF [x]) defs_sym; | 
| 947 | val t_d0 = MetaSimplifier.rewrite_rule defs_sym_eq t_l; | |
| 948 | val t_d = repeat_eqsubst_thm lthy defs_sym t_d0; | |
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 949 | val t_r = MetaSimplifier.rewrite_rule [reps_same] t_d; | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 950 | in | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 951 | ObjectLogic.rulify t_r | 
| 187 | 952 | end | 
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 953 | *} | 
| 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 954 | |
| 239 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 955 | |
| 198 
ff4425e000db
Completely cleaned Int.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
197diff
changeset | 956 | end | 
| 239 
02b14a21761a
Cleaning of the interface to lift.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
236diff
changeset | 957 |