| author | haftmann | 
| Wed, 27 Dec 2006 16:18:07 +0100 | |
| changeset 21902 | 8e5e2571c716 | 
| parent 21878 | cfc07819bb47 | 
| child 22199 | b617ddd200eb | 
| permissions | -rw-r--r-- | 
| 20319 | 1 | (* | 
| 2 | ID: $Id$ | |
| 3 | Author: Amine Chaieb, TU Muenchen | |
| 4 | ||
| 5 | A trial for automatical reification. | |
| 6 | *) | |
| 7 | ||
| 8 | signature REFLECTION = sig | |
| 9 | val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic | |
| 10 | val reflection_tac: Proof.context -> thm list -> term option -> int -> tactic | |
| 21878 | 11 | val gen_reflection_tac: Proof.context -> (cterm -> thm) | 
| 12 | -> thm list -> term option -> int -> tactic | |
| 20319 | 13 | end; | 
| 14 | ||
| 15 | structure Reflection : REFLECTION | |
| 16 | = struct | |
| 17 | ||
| 20374 | 18 | val ext2 = thm "ext2"; | 
| 21669 | 19 | val nth_Cons_0 = thm "nth_Cons_0"; | 
| 20 | val nth_Cons_Suc = thm "nth_Cons_Suc"; | |
| 21 | ||
| 20374 | 22 | (* Make a congruence rule out of a defining equation for the interpretation *) | 
| 23 | (* th is one defining equation of f, i.e. | |
| 24 | th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *) | |
| 25 | (* Cp is a constructor pattern and P is a pattern *) | |
| 26 | ||
| 27 | (* The result is: | |
| 28 | [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *) | |
| 29 | (* + the a list of names of the A1 .. An, Those are fresh in the ctxt*) | |
| 30 | ||
| 31 | ||
| 32 | fun mk_congeq ctxt fs th = | |
| 33 | let | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 34 | val (f as Const(fN,fT)) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 35 | |> fst |> strip_comb |> fst | 
| 20374 | 36 | val thy = ProofContext.theory_of ctxt | 
| 37 | val cert = Thm.cterm_of thy | |
| 38 | val (((_,_),[th']), ctxt') = Variable.import true [th] ctxt | |
| 39 | val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th')) | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 40 | fun add_fterms (t as t1 $ t2) = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 41 | if exists (fn f => (t |> strip_comb |> fst) aconv f) fs then insert (op aconv) t | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 42 | else add_fterms t1 #> add_fterms t2 | 
| 20374 | 43 | | add_fterms (t as Abs(xn,xT,t')) = | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 44 | if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => []) | 
| 20374 | 45 | | add_fterms _ = I | 
| 46 | val fterms = add_fterms rhs [] | |
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 47 | val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt' | 
| 20374 | 48 | val tys = map fastype_of fterms | 
| 49 | val vs = map Free (xs ~~ tys) | |
| 50 | val env = fterms ~~ vs | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 51 | (* FIXME!!!!*) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 52 | fun replace_fterms (t as t1 $ t2) = | 
| 20374 | 53 | (case AList.lookup (op aconv) env t of | 
| 54 | SOME v => v | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 55 | | NONE => replace_fterms t1 $ replace_fterms t2) | 
| 20374 | 56 | | replace_fterms t = (case AList.lookup (op aconv) env t of | 
| 57 | SOME v => v | |
| 58 | | NONE => t) | |
| 59 | ||
| 60 | fun mk_def (Abs(x,xT,t),v) = HOLogic.mk_Trueprop ((HOLogic.all_const xT)$ Abs(x,xT,HOLogic.mk_eq(v$(Bound 0), t))) | |
| 61 | | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t)) | |
| 62 | fun tryext x = (x RS ext2 handle _ => x) | |
| 63 | val cong = (Goal.prove ctxt'' [] (map mk_def env) | |
| 64 | (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs))) | |
| 65 | (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) | |
| 66 | THEN rtac th' 1)) RS sym | |
| 67 | ||
| 68 | val (cong' :: vars') = | |
| 69 | Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs) | |
| 70 | val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars' | |
| 71 | ||
| 72 | in (vs', cong') end; | |
| 20319 | 73 | (* congs is a list of pairs (P,th) where th is a theorem for *) | 
| 74 | (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *) | |
| 75 | val FWD = curry (op OF); | |
| 76 | ||
| 77 | (* da is the decomposition for atoms, ie. it returns ([],g) where g | |
| 78 | returns the right instance f (AtC n) = t , where AtC is the Atoms | |
| 79 | constructor and n is the number of the atom corresponding to t *) | |
| 80 | ||
| 81 | (* Generic decomp for reification : matches the actual term with the | |
| 82 | rhs of one cong rule. The result of the matching guides the | |
| 83 | proof synthesis: The matches of the introduced Variables A1 .. An are | |
| 84 | processed recursively | |
| 85 | The rest is instantiated in the cong rule,i.e. no reification is needed *) | |
| 86 | ||
| 20374 | 87 | exception REIF of string; | 
| 88 | ||
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 89 | val bds = ref ([]: (typ * ((term list) * (term list))) list); | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 90 | |
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 91 | fun index_of t = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 92 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 93 | val tt = HOLogic.listT (fastype_of t) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 94 | in | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 95 | (case AList.lookup (op =) (!bds) tt of | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 96 | NONE => error "index_of : type not found in environements!" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 97 | | SOME (tbs,tats) => | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 98 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 99 | val i = find_index_eq t tats | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 100 | val j = find_index_eq t tbs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 101 | in (if j= ~1 then | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 102 | if i= ~1 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 103 | then (bds := AList.update (op =) (tt,(tbs,tats@[t])) (!bds) ; | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 104 | length tbs + length tats) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 105 | else i else j) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 106 | end) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 107 | end; | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 108 | |
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 109 | fun dest_listT (Type ("List.list", [T])) = T;
 | 
| 20374 | 110 | |
| 111 | fun decomp_genreif da cgns (t,ctxt) = | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 112 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 113 | val thy = ProofContext.theory_of ctxt | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 114 | val cert = cterm_of thy | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 115 | fun tryabsdecomp (s,ctxt) = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 116 | (case s of | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 117 | Abs(xn,xT,ta) => | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 118 | (let | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 119 | val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 120 | val (xn,ta) = variant_abs (xn,xT,ta) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 121 | val x = Free(xn,xT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 122 | val _ = (case AList.lookup (op =) (!bds) (HOLogic.listT xT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 123 | of NONE => error "tryabsdecomp: Type not found in the Environement" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 124 | | SOME (bsT,atsT) => | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 125 | (bds := AList.update (op =) (HOLogic.listT xT, ((x::bsT), atsT)) (!bds))) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 126 | in ([(ta, ctxt')] , | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 127 | fn [th] => ((let val (bsT,asT) = the(AList.lookup (op =) (!bds) (HOLogic.listT xT)) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 128 | in (bds := AList.update (op =) (HOLogic.listT xT,(tl bsT,asT)) (!bds)) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 129 | end) ; | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 130 | hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI]))) | 
| 20374 | 131 | end) | 
| 132 | | _ => da (s,ctxt)) | |
| 133 | in | |
| 134 | (case cgns of | |
| 135 | [] => tryabsdecomp (t,ctxt) | |
| 136 | | ((vns,cong)::congs) => ((let | |
| 137 | val cert = cterm_of thy | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 138 | val certy = ctyp_of thy | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 139 | val (tyenv, tmenv) = | 
| 20319 | 140 | Pattern.match thy | 
| 141 | ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t) | |
| 142 | (Envir.type_env (Envir.empty 0),Term.Vartab.empty) | |
| 143 | val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv) | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 144 | val (fts,its) = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 145 | (map (snd o snd) fnvs, | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 146 | map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 147 | val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 148 | in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong)) | 
| 20319 | 149 | end) | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 150 | handle MATCH => decomp_genreif da congs (t,ctxt))) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 151 | end; | 
| 20319 | 152 | (* looks for the atoms equation and instantiates it with the right number *) | 
| 20374 | 153 | |
| 154 | fun mk_decompatom eqs (t,ctxt) = | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 155 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 156 | val tT = fastype_of t | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 157 | fun isat eq = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 158 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 159 | val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 160 | in exists_Const | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 161 | (fn (n,ty) => n="List.nth" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 162 | andalso | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 163 | AList.defined (op =) (!bds) (domain_type ty)) rhs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 164 | andalso fastype_of rhs = tT | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 165 | end | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 166 | fun get_nth t = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 167 | case t of | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 168 |      Const("List.nth",_)$vs$n => (t,vs,n)
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 169 | | t1$t2 => (get_nth t1 handle REIF "get_nth" => get_nth t2) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 170 | | Abs(_,_,t') => get_nth t' | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 171 | | _ => raise REIF "get_nth" | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 172 | val ([xn,vsn],ctxt') = Variable.variant_fixes ["x","vs"] ctxt | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 173 | val thy = ProofContext.theory_of ctxt' | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 174 | val cert = cterm_of thy | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 175 | fun tryeqs [] = raise REIF "Can not find the atoms equation" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 176 | | tryeqs (eq::eqs) = (( | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 177 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 178 | val rhs = eq |> prop_of |> HOLogic.dest_Trueprop | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 179 | |> HOLogic.dest_eq |> snd | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 180 | val (nt,vs,n) = get_nth rhs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 181 | val ntT = fastype_of nt | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 182 | val ntlT = HOLogic.listT ntT | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 183 | val (bsT,asT) = the (AList.lookup (op =) (!bds) ntlT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 184 | val x = Var ((xn,0),ntT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 185 | val rhs_P = subst_free [(nt,x)] rhs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 186 | val (_, tmenv) = Pattern.match | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 187 | thy (rhs_P, t) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 188 | (Envir.type_env (Envir.empty 0),Term.Vartab.empty) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 189 | val tml = Vartab.dest tmenv | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 190 | val SOME (_,t') = AList.lookup (op =) tml (xn,0) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 191 | val cvs = | 
| 21078 | 192 | 	   cert (fold_rev (fn x => fn xs => Const("List.list.Cons", ntT --> ntlT --> ntlT)$x$xs)
 | 
| 193 | bsT (Free (vsn, ntlT))) | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 194 | val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 195 | (AList.delete (op =) (xn,0) tml) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 196 | val th = (instantiate | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 197 | ([], | 
| 21621 | 198 | [(cert vs, cvs),(cert n, t' |> index_of |> IntInf.fromInt |> HOLogic.mk_nat |> cert)] | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 199 | @cts) eq) RS sym | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 200 | in hd (Variable.export ctxt' ctxt [th]) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 201 | end) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 202 | handle MATCH => tryeqs eqs) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 203 | in ([], fn _ => tryeqs (filter isat eqs)) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 204 | end; | 
| 20374 | 205 | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 206 | (* | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 207 | fun mk_decompatom eqs (t,ctxt) = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 208 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 209 | val tT = fastype_of t | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 210 | val tlT = HOLogic.listT tT | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 211 | val (bsT,asT) = (the (AList.lookup (op =) (!bds) tlT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 212 | handle Option => error "mk_decompatom: Type not found in the env.") | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 213 |   fun isateq (_$_$(Const("List.nth",_)$vs$_)) = (fastype_of vs = tlT)
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 214 | | isateq _ = false | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 215 | in case List.find (isateq o HOLogic.dest_Trueprop o prop_of) eqs of | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 216 | NONE => raise REIF "Can not find the atoms equation" | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 217 | | SOME th => | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 218 | ([], | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 219 | fn ths => | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 220 | let | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 221 | val ([x], ctxt') = Variable.variant_fixes ["vs"] ctxt | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 222 | val cert = cterm_of (ProofContext.theory_of ctxt') | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 223 |          val (Const("List.nth",_)$(vs as Var((vsn,vsi),_))$n) = 
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 224 | (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 225 | val cvs = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 226 | 	     cert (foldr (fn (x,xs) => Const("List.list.Cons", tT --> tlT --> tlT)$x$xs) 
 | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 227 | (Free(x,tlT)) bsT) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 228 | val th' = (instantiate ([], | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 229 | [(cert vs, cvs), | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 230 | (cert n, cert (HOLogic.mk_nat(index_of t)))]) th) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 231 | RS sym | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 232 | in hd (Variable.export ctxt' ctxt [th']) end) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 233 | end; | 
| 20374 | 234 | *) | 
| 20319 | 235 | (* Generic reification procedure: *) | 
| 236 | (* creates all needed cong rules and then just uses the theorem synthesis *) | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 237 | |
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 238 | fun mk_congs ctxt raw_eqs = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 239 | let | 
| 21078 | 240 | val fs = fold_rev (fn eq => | 
| 20853 | 241 | insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 242 | |> HOLogic.dest_eq |> fst |> strip_comb | 
| 21078 | 243 | |> fst)) raw_eqs [] | 
| 244 | val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> split_last |> fst) | |
| 245 | union ts) fs [] | |
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 246 | val _ = bds := AList.make (fn _ => ([],[])) tys | 
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 247 | val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 248 | val thy = ProofContext.theory_of ctxt' | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 249 | val cert = cterm_of thy | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 250 | val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t))))) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 251 | (tys ~~ vs) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 252 | fun insteq eq ts = | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 253 | let val itms = map (fn t => t|> (AList.lookup (op =) vstys) |> the) ts | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 254 | in instantiate' [] itms eq | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 255 | end | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 256 | val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 257 | |> HOLogic.dest_eq |> fst |> strip_comb |> fst |> fastype_of | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 258 | |> binder_types |> split_last |> fst | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 259 | |> (insteq eq)) raw_eqs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 260 | val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 261 | in ps ~~ (Variable.export ctxt' ctxt congs) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 262 | end; | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 263 | |
| 20319 | 264 | fun genreif ctxt raw_eqs t = | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 265 | let | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 266 | val _ = bds := [] | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 267 | val congs = mk_congs ctxt raw_eqs | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 268 | val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 269 | val tys = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 270 | |> strip_comb |> fst |> fastype_of |> strip_type |> fst | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 271 | |> split_last |> fst | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 272 | val cert = cterm_of (ProofContext.theory_of ctxt) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 273 | val cvs = map (fn t => t |> (AList.lookup (op =) (!bds)) |> the |> snd | 
| 21757 | 274 | |> HOLogic.mk_list (dest_listT t) |> cert |> SOME) | 
| 20564 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 275 | tys | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 276 | val th' = (instantiate' [] cvs (th RS sym)) RS sym | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 277 | val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th' | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 278 | val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 279 | (fn _ => Simp_tac 1) | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 280 | val _ = bds := [] | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 281 | in FWD trans [th'',th'] | 
| 
6857bd9f1a79
Reifiaction now deals with Interpretations with an arbtrary number of parameters. It deals with binding. The Atomic cases can be I ... = f (xs!n)
 chaieb parents: 
20374diff
changeset | 282 | end; | 
| 20319 | 283 | |
| 21878 | 284 | fun genreflect ctxt conv corr_thm raw_eqs t = | 
| 20319 | 285 | let val th = FWD trans [genreif ctxt raw_eqs t, corr_thm RS sym] | 
| 286 | val ft = (snd o Thm.dest_comb o snd o Thm.dest_comb o snd o Thm.dest_comb o cprop_of) th | |
| 21878 | 287 | val rth = conv ft | 
| 20319 | 288 | in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc]) | 
| 289 | (simplify (HOL_basic_ss addsimps [rth]) th) | |
| 290 | end | |
| 291 | ||
| 292 | fun genreify_tac ctxt eqs to i = (fn st => | |
| 293 | let | |
| 294 | val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1)) | |
| 295 | val t = (case to of NONE => P | SOME x => x) | |
| 296 | val th = (genreif ctxt eqs t) RS ssubst | |
| 297 | in rtac th i st | |
| 298 | end); | |
| 299 | ||
| 300 | (* Reflection calls reification and uses the correctness *) | |
| 301 | (* theorem assumed to be the dead of the list *) | |
| 21878 | 302 | fun gen_reflection_tac ctxt conv (corr_thm :: raw_eqs) to i = (fn st => | 
| 303 | let | |
| 304 | val P = HOLogic.dest_Trueprop (nth (prems_of st) (i - 1)); | |
| 305 | val t = the_default P to; | |
| 306 | val th = genreflect ctxt conv corr_thm raw_eqs t | |
| 307 | RS ssubst; | |
| 308 | in rtac th i st end); | |
| 309 | ||
| 310 | fun reflection_tac ctxt = gen_reflection_tac ctxt NBE.normalization_conv; | |
| 20319 | 311 | |
| 20797 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 wenzelm parents: 
20595diff
changeset | 312 | end |