| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Mon, 29 Mar 2010 16:44:05 +0200 | |
| changeset 1696 | fb4be1983a71 | 
| parent 1685 | 721d92623c9d | 
| child 1774 | c34347ec7ab3 | 
| permissions | -rw-r--r-- | 
| 1274 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 1 | theory Lift | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 2 | imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" | 
| 1274 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | |
| 1681 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 5 | |
| 1316 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 6 | ML {*
 | 
| 1681 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 7 | fun define_quotient_types binds tys alphas equivps ctxt = | 
| 1656 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 8 | let | 
| 1681 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 9 | fun def_ty ((b, ty), (alpha, equivp)) ctxt = | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 10 | Quotient_Type.add_quotient_type ((([], b, NoSyn), (ty, alpha)), equivp) ctxt; | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 11 | val alpha_equivps = List.take (equivps, length alphas) | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 12 | val (thms, ctxt') = fold_map def_ty ((binds ~~ tys) ~~ (alphas ~~ alpha_equivps)) ctxt; | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 13 | val quot_thms = map fst thms; | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 14 | val quots = map (HOLogic.dest_Trueprop o prop_of) quot_thms; | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 15 | val reps = map (hd o rev o snd o strip_comb) quots; | 
| 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 16 | val qtys = map (domain_type o fastype_of) reps; | 
| 1656 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 17 | in | 
| 1681 
b8a07a3c1692
Get lifted types information from the quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1656diff
changeset | 18 | (qtys, ctxt') | 
| 1656 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 19 | end | 
| 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 20 | *} | 
| 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 21 | |
| 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 22 | (* Renames schematic variables in a theorem *) | 
| 
c9d3dda79fe3
Removed remaining cheats + some cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1553diff
changeset | 23 | ML {*
 | 
| 1498 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 24 | fun rename_vars fnctn thm = | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 25 | let | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 26 | val vars = Term.add_vars (prop_of thm) [] | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 27 | val nvars = map (Var o ((apfst o apfst) fnctn)) vars | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 28 | in | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 29 | Thm.certify_instantiate ([], (vars ~~ nvars)) thm | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 30 | end | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 31 | *} | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 32 | |
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 33 | ML {*
 | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 34 | fun un_raws name = | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 35 | let | 
| 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 36 | fun un_raw name = unprefix "_raw" name handle Fail _ => name | 
| 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 37 | fun add_under names = hd names :: (map (prefix "_") (tl names)) | 
| 1498 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 38 | in | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 39 | implode (map un_raw (add_under (space_explode "_" name))) | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 40 | end | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 41 | *} | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 42 | |
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 43 | (* Similar to Tools/IsaPlanner/rw_tools.ML *) | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 44 | ML {*
 | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 45 | fun rename_term_bvars (Abs(s, ty, t)) = (Abs(un_raws s, ty, rename_term_bvars t)) | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 46 | | rename_term_bvars (a $ b) = (rename_term_bvars a) $ (rename_term_bvars b) | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 47 | | rename_term_bvars x = x; | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 48 | |
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 49 | fun rename_thm_bvars th = | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 50 | let | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 51 | val t = Thm.prop_of th | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 52 | in | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 53 | Thm.rename_boundvars t (rename_term_bvars t) th | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 54 | end; | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 55 | *} | 
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 56 | |
| 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 57 | ML {*
 | 
| 1683 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1681diff
changeset | 58 | fun lift_thm qtys ctxt thm = | 
| 1498 
2ff84b1f551f
Rename bound variables + minor cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1497diff
changeset | 59 | let | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 60 | val un_raw_names = rename_vars un_raws | 
| 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 61 | in | 
| 1683 
f78c820f67c3
Automatically lift theorems and constants only using the new quotient types. Requires new Isabelle.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1681diff
changeset | 62 | rename_thm_bvars (un_raw_names (Quotient_Tacs.lifted qtys ctxt thm)) | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 63 | end | 
| 1276 
3365fce80f0f
To call quotient it is enough to export the alpha frees to proper constants and their respective equivp theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1274diff
changeset | 64 | *} | 
| 1274 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 65 | |
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 66 | ML {*
 | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 67 | fun define_fv_alpha_export dt binds bns ctxt = | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 68 | let | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1498diff
changeset | 69 | val ((((fv_ts_loc, fv_def_loc), ord_fv_ts_loc), alpha), ctxt') = | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 70 | define_fv_alpha dt binds bns ctxt; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 71 | val alpha_ts_loc = #preds alpha | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 72 | val alpha_induct_loc = #induct alpha | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 73 | val alpha_intros_loc = #intrs alpha; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 74 | val alpha_cases_loc = #elims alpha | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 75 | val morphism = ProofContext.export_morphism ctxt' ctxt; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 76 | val fv_ts = map (Morphism.term morphism) fv_ts_loc; | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1498diff
changeset | 77 | val ord_fv_ts = map (Morphism.term morphism) ord_fv_ts_loc; | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 78 | val fv_def = Morphism.fact morphism fv_def_loc; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 79 | val alpha_ts = map (Morphism.term morphism) alpha_ts_loc; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 80 | val alpha_induct = Morphism.thm morphism alpha_induct_loc; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 81 | val alpha_intros = Morphism.fact morphism alpha_intros_loc | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 82 | val alpha_cases = Morphism.fact morphism alpha_cases_loc | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 83 | in | 
| 1553 
4355eb3b7161
Automatically derive support for datatypes with at-most one binding per constructor.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1498diff
changeset | 84 | ((((fv_ts, ord_fv_ts), fv_def), ((alpha_ts, alpha_intros), (alpha_cases, alpha_induct))), ctxt') | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 85 | end; | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 86 | *} | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 87 | |
| 1280 
1f057f8da8aa
Progress with general lifting procedure.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1277diff
changeset | 88 | end | 
| 1274 
d867021d8ac1
Preparing the generalized lifting procedure
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 |