src/HOL/ex/reflection.ML
author haftmann
Tue, 19 Sep 2006 15:21:58 +0200
changeset 20595 db6bedfba498
parent 20564 6857bd9f1a79
child 20797 c1f0bc7e7d80
permissions -rw-r--r--
improved numeral handling for nbe
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     1
(*
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     4
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     5
A trial for automatical reification.
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     6
*)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     7
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     8
signature REFLECTION = sig
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     9
  val genreify_tac: Proof.context -> thm list -> term option -> int -> tactic
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    10
  val reflection_tac: Proof.context -> thm list -> term option -> int -> tactic
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    11
end;
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    12
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    13
structure Reflection : REFLECTION
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    14
= struct
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    15
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    16
val ext2 = thm "ext2";
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    17
  (* Make a congruence rule out of a defining equation for the interpretation *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    18
  (* th is one defining equation of f, i.e.
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    19
     th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    20
  (* Cp is a constructor pattern and P is a pattern *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    21
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    22
  (* The result is:
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    23
      [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    24
  (*  + the a list of names of the A1 .. An, Those are fresh in the ctxt*)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    25
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    26
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    27
fun mk_congeq ctxt fs th = 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    28
  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: 20374
diff changeset
    29
   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: 20374
diff changeset
    30
				|> fst |> strip_comb |> fst
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    31
   val thy = ProofContext.theory_of ctxt
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    32
   val cert = Thm.cterm_of thy
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    33
   val (((_,_),[th']), ctxt') = Variable.import true [th] ctxt
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    34
   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: 20374
diff changeset
    35
   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: 20374
diff changeset
    36
       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: 20374
diff changeset
    37
       else add_fterms t1 #> add_fterms t2
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    38
     | 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: 20374
diff changeset
    39
       if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => [])
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    40
     | add_fterms _ = I
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    41
   val fterms = add_fterms rhs []
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    42
   val (xs, ctxt'') = Variable.invent_fixes (replicate (length fterms) "x") ctxt'
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    43
   val tys = map fastype_of fterms
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    44
   val vs = map Free (xs ~~ tys)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    45
   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: 20374
diff changeset
    46
		    (* 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: 20374
diff changeset
    47
   fun replace_fterms (t as t1 $ t2) =
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    48
       (case AList.lookup (op aconv) env t of
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    49
	    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: 20374
diff changeset
    50
	  | NONE => replace_fterms t1 $ replace_fterms t2)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    51
     | replace_fterms t = (case AList.lookup (op aconv) env t of
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    52
			       SOME v => v
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    53
			     | NONE => t)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    54
      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    55
   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)))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    56
     | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    57
   fun tryext x = (x RS ext2 handle _ =>  x)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    58
   val cong = (Goal.prove ctxt'' [] (map mk_def env)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    59
			  (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    60
			  (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    61
							THEN rtac th' 1)) RS sym
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    62
	      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    63
   val (cong' :: vars') = 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    64
       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    65
   val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    66
					      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    67
  in  (vs', cong') end; 
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    68
 (* congs is a list of pairs (P,th) where th is a theorem for *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    69
        (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    70
val FWD = curry (op OF);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    71
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    72
 (* da is the decomposition for atoms, ie. it returns ([],g) where g
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    73
 returns the right instance f (AtC n) = t , where AtC is the Atoms
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    74
 constructor and n is the number of the atom corresponding to t *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    75
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    76
(* Generic decomp for reification : matches the actual term with the
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    77
rhs of one cong rule. The result of the matching guides the
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    78
proof synthesis: The matches of the introduced Variables A1 .. An are
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    79
processed recursively
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    80
 The rest is instantiated in the cong rule,i.e. no reification is needed *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    81
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    82
exception REIF of string;
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    83
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: 20374
diff changeset
    84
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: 20374
diff changeset
    85
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: 20374
diff changeset
    86
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: 20374
diff changeset
    87
 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: 20374
diff changeset
    88
  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: 20374
diff changeset
    89
 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: 20374
diff changeset
    90
  (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: 20374
diff changeset
    91
    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: 20374
diff changeset
    92
  | 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: 20374
diff changeset
    93
    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: 20374
diff changeset
    94
     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: 20374
diff changeset
    95
     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: 20374
diff changeset
    96
    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: 20374
diff changeset
    97
	    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: 20374
diff changeset
    98
	    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: 20374
diff changeset
    99
		  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: 20374
diff changeset
   100
	    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: 20374
diff changeset
   101
    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: 20374
diff changeset
   102
 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: 20374
diff changeset
   103
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: 20374
diff changeset
   104
fun dest_listT (Type ("List.list", [T])) = T;
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   105
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   106
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: 20374
diff changeset
   107
 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: 20374
diff changeset
   108
  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: 20374
diff changeset
   109
  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: 20374
diff changeset
   110
  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: 20374
diff changeset
   111
   (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: 20374
diff changeset
   112
     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: 20374
diff changeset
   113
     (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: 20374
diff changeset
   114
       val ([xn],ctxt') = Variable.invent_fixes ["x"] 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: 20374
diff changeset
   115
       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: 20374
diff changeset
   116
       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: 20374
diff changeset
   117
       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: 20374
diff changeset
   118
		 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: 20374
diff changeset
   119
		  | 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: 20374
diff changeset
   120
		    (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: 20374
diff changeset
   121
      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: 20374
diff changeset
   122
	  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: 20374
diff changeset
   123
		       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: 20374
diff changeset
   124
		       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: 20374
diff changeset
   125
		      hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI])))
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   126
	end)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   127
    | _ => da (s,ctxt))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   128
  in 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   129
  (case cgns of 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   130
    [] => tryabsdecomp (t,ctxt)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   131
  | ((vns,cong)::congs) => ((let
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   132
        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: 20374
diff changeset
   133
	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: 20374
diff changeset
   134
        val (tyenv, tmenv) =
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   135
        Pattern.match thy
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   136
        ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   137
        (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   138
        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: 20374
diff changeset
   139
        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: 20374
diff changeset
   140
	    (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: 20374
diff changeset
   141
             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: 20374
diff changeset
   142
	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: 20374
diff changeset
   143
    in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   144
    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: 20374
diff changeset
   145
      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: 20374
diff changeset
   146
  end;
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   147
          (* looks for the atoms equation and instantiates it with the right number *)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   148
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   149
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: 20374
diff changeset
   150
 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: 20374
diff changeset
   151
  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: 20374
diff changeset
   152
  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: 20374
diff changeset
   153
   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: 20374
diff changeset
   154
    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: 20374
diff changeset
   155
   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: 20374
diff changeset
   156
	  (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: 20374
diff changeset
   157
			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: 20374
diff changeset
   158
			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: 20374
diff changeset
   159
	  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: 20374
diff changeset
   160
   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: 20374
diff changeset
   161
  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: 20374
diff changeset
   162
   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: 20374
diff changeset
   163
     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: 20374
diff changeset
   164
   | 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: 20374
diff changeset
   165
   | 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: 20374
diff changeset
   166
   | _ => raise REIF "get_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: 20374
diff changeset
   167
  val ([xn,vsn],ctxt') = Variable.invent_fixes ["x","vs"] 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: 20374
diff changeset
   168
  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: 20374
diff changeset
   169
  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: 20374
diff changeset
   170
  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: 20374
diff changeset
   171
    | 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: 20374
diff changeset
   172
      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: 20374
diff changeset
   173
       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: 20374
diff changeset
   174
			  |> 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: 20374
diff changeset
   175
       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: 20374
diff changeset
   176
       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: 20374
diff changeset
   177
       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: 20374
diff changeset
   178
       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: 20374
diff changeset
   179
       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: 20374
diff changeset
   180
       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: 20374
diff changeset
   181
       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: 20374
diff changeset
   182
			    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: 20374
diff changeset
   183
			    (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: 20374
diff changeset
   184
       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: 20374
diff changeset
   185
       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: 20374
diff changeset
   186
       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: 20374
diff changeset
   187
	   cert (foldr (fn (x,xs) => Const("List.list.Cons", ntT --> ntlT --> ntlT)$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: 20374
diff changeset
   188
		       (Free(vsn,ntlT)) 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: 20374
diff changeset
   189
       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: 20374
diff changeset
   190
		     (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: 20374
diff changeset
   191
       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: 20374
diff changeset
   192
		     ([],
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: 20374
diff changeset
   193
		      [(cert vs, cvs),(cert n, t' |> index_of |> HOLogic.mk_nat |> cert)]
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: 20374
diff changeset
   194
		      @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: 20374
diff changeset
   195
      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: 20374
diff changeset
   196
      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: 20374
diff changeset
   197
	  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: 20374
diff changeset
   198
 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: 20374
diff changeset
   199
 end;
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   200
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: 20374
diff changeset
   201
(* 
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: 20374
diff changeset
   202
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: 20374
diff changeset
   203
 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: 20374
diff changeset
   204
  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: 20374
diff changeset
   205
  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: 20374
diff changeset
   206
  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: 20374
diff changeset
   207
		   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: 20374
diff changeset
   208
  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: 20374
diff changeset
   209
    | 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: 20374
diff changeset
   210
 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: 20374
diff changeset
   211
     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: 20374
diff changeset
   212
   | 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: 20374
diff changeset
   213
     ([],
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: 20374
diff changeset
   214
      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: 20374
diff changeset
   215
        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: 20374
diff changeset
   216
         val ([x], ctxt') = Variable.invent_fixes ["vs"] 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: 20374
diff changeset
   217
	 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: 20374
diff changeset
   218
         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: 20374
diff changeset
   219
             (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: 20374
diff changeset
   220
         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: 20374
diff changeset
   221
	     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: 20374
diff changeset
   222
			 (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: 20374
diff changeset
   223
	 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: 20374
diff changeset
   224
				 [(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: 20374
diff changeset
   225
				  (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: 20374
diff changeset
   226
		       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: 20374
diff changeset
   227
        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: 20374
diff changeset
   228
 end;
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   229
*)
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   230
  (* Generic reification procedure: *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   231
  (* 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: 20374
diff changeset
   232
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: 20374
diff changeset
   233
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: 20374
diff changeset
   234
 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: 20374
diff changeset
   235
  val fs = foldr (fn (eq,fns) => 
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: 20374
diff changeset
   236
		     (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: 20374
diff changeset
   237
			 |> HOLogic.dest_eq |> fst |> strip_comb 
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: 20374
diff changeset
   238
			 |> fst) ins fns) [] 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: 20374
diff changeset
   239
  val tys = foldr (fn (f,ts) => (f |> fastype_of |> 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: 20374
diff changeset
   240
				    union ts) [] fs
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: 20374
diff changeset
   241
  val _ = bds := AList.make (fn _ => ([],[])) 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: 20374
diff changeset
   242
  val (vs, ctxt') = Variable.invent_fixes (replicate (length tys) "vs") 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: 20374
diff changeset
   243
  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: 20374
diff changeset
   244
  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: 20374
diff changeset
   245
  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: 20374
diff changeset
   246
		  (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: 20374
diff changeset
   247
  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: 20374
diff changeset
   248
   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: 20374
diff changeset
   249
   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: 20374
diff changeset
   250
   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: 20374
diff changeset
   251
  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: 20374
diff changeset
   252
			     |> 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: 20374
diff changeset
   253
			     |> 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: 20374
diff changeset
   254
			     |> (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: 20374
diff changeset
   255
  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: 20374
diff changeset
   256
 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: 20374
diff changeset
   257
 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: 20374
diff changeset
   258
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   259
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: 20374
diff changeset
   260
 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: 20374
diff changeset
   261
  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: 20374
diff changeset
   262
  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: 20374
diff changeset
   263
  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: 20374
diff changeset
   264
  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: 20374
diff changeset
   265
	       |> 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: 20374
diff changeset
   266
	       |> 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: 20374
diff changeset
   267
  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: 20374
diff changeset
   268
  val cvs = map (fn t => t |> (AList.lookup (op =) (!bds)) |> the |> 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: 20374
diff changeset
   269
			   |> HOLogic.mk_list I (dest_listT t) |> cert |> SOME)
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: 20374
diff changeset
   270
		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: 20374
diff changeset
   271
  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: 20374
diff changeset
   272
  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: 20374
diff changeset
   273
  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: 20374
diff changeset
   274
			(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: 20374
diff changeset
   275
  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: 20374
diff changeset
   276
 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: 20374
diff changeset
   277
 end;
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   278
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   279
fun genreflect ctxt corr_thm raw_eqs t =
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   280
    let val th = FWD trans [genreif ctxt raw_eqs t, corr_thm RS sym]
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   281
        val ft = (snd o Thm.dest_comb o snd o Thm.dest_comb o snd o Thm.dest_comb o cprop_of) th
20595
db6bedfba498 improved numeral handling for nbe
haftmann
parents: 20564
diff changeset
   282
        val rth = NBE.normalization_conv ft
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   283
    in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc])
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   284
                (simplify (HOL_basic_ss addsimps [rth]) th)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   285
    end
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   286
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   287
fun genreify_tac ctxt eqs to i = (fn st =>
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   288
  let
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   289
    val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   290
    val t = (case to of NONE => P | SOME x => x)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   291
    val th = (genreif ctxt eqs t) RS ssubst
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   292
  in rtac th i st
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   293
  end);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   294
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   295
    (* Reflection calls reification and uses the correctness *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   296
        (* theorem assumed to be the dead of the list *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   297
 fun reflection_tac ctxt (corr_thm::raw_eqs) to i =
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   298
    (fn st =>
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   299
        let val P = (HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1)))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   300
            val t = (case to of NONE => P | SOME x => x)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   301
            val th = (genreflect ctxt corr_thm raw_eqs t) RS ssubst
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   302
        in rtac th i st end);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   303
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: 20374
diff changeset
   304
end