src/HOL/ex/reflection.ML
author wenzelm
Tue, 18 Mar 2008 21:57:36 +0100
changeset 26321 d875e70a94de
parent 24630 351a308ab58d
child 29269 5c25a2012975
permissions -rw-r--r--
valid_thms: get_thms_silent;
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
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
    10
  val reflection_tac: Proof.context -> thm list -> thm list -> term option -> int -> tactic
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
    11
  val gen_reflection_tac: Proof.context -> (cterm -> thm)
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
    12
    -> thm list -> thm list -> term option -> int -> tactic
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    13
end;
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    14
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
    15
structure Reflection : REFLECTION
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    16
= struct
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    17
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    18
val ext2 = thm "ext2";
21669
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21621
diff changeset
    19
val nth_Cons_0 = thm "nth_Cons_0";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21621
diff changeset
    20
val nth_Cons_Suc = thm "nth_Cons_Suc";
c68717c16013 removed legacy ML bindings;
wenzelm
parents: 21621
diff changeset
    21
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    22
  (* Make a congruence rule out of a defining equation for the interpretation *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    23
  (* th is one defining equation of f, i.e.
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    24
     th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    25
  (* Cp is a constructor pattern and P is a pattern *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    26
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    27
  (* The result is:
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    28
      [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    29
  (*  + 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
    30
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    31
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    32
fun mk_congeq ctxt fs th = 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff 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: 20374
diff changeset
    35
				|> fst |> strip_comb |> fst
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    36
   val thy = ProofContext.theory_of ctxt
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    37
   val cert = Thm.cterm_of thy
22568
ed7aa5a350ef renamed Variable.import to import_thms (avoid clash with Alice keywords);
wenzelm
parents: 22199
diff changeset
    38
   val (((_,_),[th']), ctxt') = Variable.import_thms true [th] ctxt
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff changeset
    40
   fun add_fterms (t as t1 $ t2) = 
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 21878
diff changeset
    41
       if exists (fn f => could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) 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
    42
       else add_fterms t1 #> add_fterms t2
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff changeset
    44
       if (fN mem (term_consts t)) then (fn _ => [t]) else (fn _ => [])
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    45
     | add_fterms _ = I
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    46
   val fterms = add_fterms rhs []
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff changeset
    47
   val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    48
   val tys = map fastype_of fterms
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    49
   val vs = map Free (xs ~~ tys)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff 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: 20374
diff changeset
    52
   fun replace_fterms (t as t1 $ t2) =
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    53
       (case AList.lookup (op aconv) env t of
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff changeset
    55
	  | NONE => replace_fterms t1 $ replace_fterms t2)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    56
     | replace_fterms t = (case AList.lookup (op aconv) env t of
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    57
			       SOME v => v
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    58
			     | NONE => t)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    59
      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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)))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    61
     | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
    62
   fun tryext x = (x RS ext2 handle THM _ =>  x)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    63
   val cong = (Goal.prove ctxt'' [] (map mk_def env)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    64
			  (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    65
			  (fn x => LocalDefs.unfold_tac (#context x) (map tryext (#prems x)) 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    66
							THEN rtac th' 1)) RS sym
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    67
	      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    68
   val (cong' :: vars') = 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    69
       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    70
   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
    71
					      
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    72
  in  (vs', cong') end; 
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    73
 (* 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
    74
        (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    75
val FWD = curry (op OF);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    76
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    77
 (* da is the decomposition for atoms, ie. it returns ([],g) where g
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    78
 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
    79
 constructor and n is the number of the atom corresponding to t *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    80
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    81
(* Generic decomp for reification : matches the actual term with the
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    82
rhs of one cong rule. The result of the matching guides the
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    83
proof synthesis: The matches of the introduced Variables A1 .. An are
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    84
processed recursively
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    85
 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
    86
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    87
exception REIF of string;
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff changeset
    94
 in 
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 21878
diff changeset
    95
  (case AList.lookup Type.could_unify (!bds) tt of
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
    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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff changeset
   102
	    if i= ~1 
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   103
	    then (bds := AList.update Type.could_unify (tt,(tbs,tats@[t])) (!bds) ; 
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
   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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff changeset
   109
fun dest_listT (Type ("List.list", [T])) = T;
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   110
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff changeset
   118
     (let
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff 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: 20374
diff 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: 20374
diff changeset
   121
       val x = Free(xn,xT)
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   122
       val _ = (case AList.lookup Type.could_unify (!bds) (HOLogic.listT xT)
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
   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: 20374
diff changeset
   124
		  | SOME (bsT,atsT) => 
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   125
		    (bds := AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) (!bds)))
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
   126
      in ([(ta, ctxt')] , 
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   127
	  fn [th] => ((let val (bsT,asT) = the(AList.lookup Type.could_unify (!bds) (HOLogic.listT xT))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   128
		       in (bds := AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) (!bds))
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
   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: 20374
diff changeset
   130
		      hd (Variable.export ctxt' ctxt [(forall_intr (cert x) th) COMP allI])))
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   131
	end)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   132
    | _ => da (s,ctxt))
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   133
  in 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   134
  (case cgns of 
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   135
    [] => tryabsdecomp (t,ctxt)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   136
  | ((vns,cong)::congs) => ((let
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   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: 20374
diff 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: 20374
diff changeset
   139
        val (tyenv, tmenv) =
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   140
        Pattern.match thy
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   141
        ((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
   142
        (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff 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: 20374
diff changeset
   148
    in (fts ~~ (replicate (length fts) ctxt), FWD (instantiate (ctyenv, its) cong))
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   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: 20374
diff 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: 20374
diff changeset
   151
  end;
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   152
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   153
 (* looks for the atoms equation and instantiates it with the right number *)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   154
23548
e25991f126ce Generalized case for atoms. Selection of environment lists is allowed more than once.
chaieb
parents: 22568
diff changeset
   155
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   156
fun mk_decompatom eqs (t,ctxt) =
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   157
let 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   158
 val tT = fastype_of t
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   159
 fun isat eq = 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   160
  let 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   161
   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
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
   162
   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
   163
	  (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
   164
			andalso 
22199
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 21878
diff changeset
   165
			AList.defined Type.could_unify (!bds) (domain_type ty)) rhs 
b617ddd200eb Now deals with simples cases where the input equations contain type variables
chaieb
parents: 21878
diff changeset
   166
	  andalso Type.could_unify (fastype_of rhs, tT)
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
   167
   end
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   168
 fun get_nths t acc = 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   169
  case t of
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   170
    Const("List.nth",_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   171
  | t1$t2 => get_nths t1 (get_nths t2 acc)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   172
  | Abs(_,_,t') => get_nths t'  acc
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   173
  | _ => acc
23548
e25991f126ce Generalized case for atoms. Selection of environment lists is allowed more than once.
chaieb
parents: 22568
diff changeset
   174
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   175
 fun 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   176
   tryeqs [] = error "Can not find the atoms equation"
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   177
 | tryeqs (eq::eqs) = ((
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   178
  let 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   179
   val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   180
   val nths = get_nths rhs []
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   181
   val (vss,ns) = fold_rev (fn (_,(vs,n)) => fn (vss,ns) => 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   182
                             (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([],[]) 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   183
   val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   184
   val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt' 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   185
   val thy = ProofContext.theory_of ctxt''
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   186
   val cert = cterm_of thy
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   187
   val certT = ctyp_of thy
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   188
   val vsns_map = vss ~~ vsns
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   189
   val xns_map = (fst (split_list nths)) ~~ xns
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   190
   val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   191
   val rhs_P = subst_free subst rhs
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   192
   val (tyenv, tmenv) = Pattern.match 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   193
	                    thy (rhs_P, t)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   194
	                    (Envir.type_env (Envir.empty 0),Term.Vartab.empty)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   195
   val sbst = Envir.subst_vars (tyenv, tmenv)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   196
   val sbsT = Envir.typ_subst_TVars tyenv
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   197
   val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   198
                      (Vartab.dest tyenv)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   199
   val tml = Vartab.dest tmenv
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   200
   val t's = map (fn xn => snd (valOf (AList.lookup (op =) tml (xn,0)))) xns (* FIXME : Express with sbst*)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   201
   val subst_ns = map (fn (Const _ $ vs $ n, Var (xn0,T)) => 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   202
                          (cert n, snd (valOf (AList.lookup (op =) tml xn0)) 
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 24044
diff changeset
   203
                             |> (index_of #> HOLogic.mk_nat #> cert))) 
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   204
                      subst
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   205
   val subst_vs = 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   206
    let 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   207
     fun ty (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,T)) = (certT T, certT (sbsT T))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   208
     fun h (Const _ $ (vs as Var (vsn,lT)) $ n, Var (xn0,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
   209
      let 
23605
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   210
       val cns = sbst (Const("List.list.Cons", T --> lT --> lT))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   211
       val lT' = sbsT lT
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   212
       val (bsT,asT) = the (AList.lookup Type.could_unify (!bds) lT)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   213
       val vsn = valOf (AList.lookup (op =) vsns_map vs)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   214
       val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   215
      in (cert vs, cvs) end
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   216
    in map h subst end
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   217
   val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   218
                 (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b)) 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   219
                       (map (fn n => (n,0)) xns) tml)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   220
   val substt = 
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   221
    let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   222
    in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   223
   val th = (instantiate (subst_ty, substt)  eq) RS sym
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   224
  in  hd (Variable.export ctxt'' ctxt [th]) end)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   225
 handle MATCH => tryeqs eqs)
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   226
in ([], fn _ => tryeqs (filter isat eqs))
81d0fdec9edc Reification now deals with type variables
chaieb
parents: 23548
diff changeset
   227
end;
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   228
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   229
  (* Generic reification procedure: *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   230
  (* 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
   231
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   232
  fun mk_congs ctxt raw_eqs = 
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   233
 let
21078
101aefd61aac slight cleanup
haftmann
parents: 20853
diff changeset
   234
  val fs = fold_rev (fn eq =>
20853
3ff5a2e05810 *** empty log message ***
haftmann
parents: 20797
diff changeset
   235
		     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: 20374
diff changeset
   236
			 |> HOLogic.dest_eq |> fst |> strip_comb 
21078
101aefd61aac slight cleanup
haftmann
parents: 20853
diff changeset
   237
			 |> fst)) raw_eqs []
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   238
  val tys = fold_rev (fn f => fn ts => (f |> fastype_of |> binder_types |> tl) 
21078
101aefd61aac slight cleanup
haftmann
parents: 20853
diff changeset
   239
				    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: 20374
diff changeset
   240
  val _ = bds := AList.make (fn _ => ([],[])) tys
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff changeset
   241
  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: 20374
diff changeset
   242
  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
   243
  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
   244
  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
   245
		  (tys ~~ vs)
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   246
  val is_Var = can dest_Var
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   247
  fun insteq eq vs = 
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   248
   let
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   249
     val subst = map (fn (v as Var(n,t)) => (cert v, (valOf o valOf) (AList.lookup (op =) vstys t)))  
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   250
  (filter is_Var vs)
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   251
   in Thm.instantiate ([],subst) eq
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
   252
   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
   253
  val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop 
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   254
			     |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
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
   255
			     |> (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
   256
  val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   257
in ps ~~ (Variable.export ctxt' ctxt congs)
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   258
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
   259
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   260
fun partition P [] = ([],[])
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   261
  | partition P (x::xs) = 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   262
     let val (yes,no) = partition P xs
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   263
     in if P x then (x::yes,no) else (yes, x::no) end
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   264
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   265
fun rearrange congs = 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   266
let 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   267
 fun P (_, th) = 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   268
  let val @{term "Trueprop"}$(Const ("op =",_) $l$_) = concl_of th
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   269
  in can dest_Var l end
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   270
 val (yes,no) = partition P congs 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   271
 in no @ yes end
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   272
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   273
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   274
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   275
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
   276
 let 
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   277
  val congs = rearrange (mk_congs ctxt raw_eqs)
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
   278
  val th = divide_and_conquer (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt)
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   279
  fun is_listVar (Var (_,t)) = can dest_listT t
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   280
       | is_listVar _ = false
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   281
  val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   282
	       |> strip_comb |> snd |> filter is_listVar
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
   283
  val cert = cterm_of (ProofContext.theory_of ctxt)
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   284
  val cvs = map (fn (v as Var(n,t)) => (cert v, the (AList.lookup Type.could_unify (!bds) t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   285
  val th' = instantiate ([], cvs) 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
   286
  val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
23643
32ee4111d1bc Corrected erronus use of compiletime context to the runtime context
chaieb
parents: 23624
diff changeset
   287
  val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
24044
8c168f5ef221 avoid ill-defined Simp_tac;
wenzelm
parents: 23791
diff changeset
   288
			(fn _ => simp_tac (local_simpset_of ctxt) 1)
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
   289
  val _ = bds := []
23624
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   290
in FWD trans [th'',th']
82091387f6d7 The order for parameter for interpretation is now inversted:
chaieb
parents: 23605
diff changeset
   291
end
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   292
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   293
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   294
fun genreflect ctxt conv corr_thms raw_eqs t =
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   295
let 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   296
  val reifth = genreif ctxt raw_eqs t
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   297
  fun trytrans [] = error "No suitable correctness theorem found"
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   298
    | trytrans (th::ths) = 
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   299
         (FWD trans [reifth, th RS sym] handle THM _ => trytrans ths)
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   300
  val th = trytrans corr_thms
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   301
  val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   302
  val rth = conv ft
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   303
in simplify (HOL_basic_ss addsimps raw_eqs addsimps [nth_Cons_0, nth_Cons_Suc])
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   304
           (simplify (HOL_basic_ss addsimps [rth]) th)
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   305
end
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   306
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   307
fun genreify_tac ctxt eqs to i = (fn st =>
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   308
  let
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   309
    val P = HOLogic.dest_Trueprop (List.nth (prems_of st, i - 1))
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   310
    val t = (case to of NONE => P | SOME x => x)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   311
    val th = (genreif ctxt eqs t) RS ssubst
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   312
  in rtac th i st
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   313
  end);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   314
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   315
    (* Reflection calls reification and uses the correctness *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   316
        (* theorem assumed to be the dead of the list *)
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   317
fun gen_reflection_tac ctxt conv corr_thms raw_eqs to i = (fn st =>
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   318
  let
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   319
    val P = HOLogic.dest_Trueprop (nth (prems_of st) (i - 1));
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   320
    val t = the_default P to;
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   321
    val th = genreflect ctxt conv corr_thms raw_eqs t
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   322
      RS ssubst;
23791
e105381d4140 tries to solve goal via TrueI
nipkow
parents: 23648
diff changeset
   323
  in (rtac th i THEN TRY(rtac TrueI i)) st end);
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   324
23791
e105381d4140 tries to solve goal via TrueI
nipkow
parents: 23648
diff changeset
   325
fun reflection_tac ctxt = gen_reflection_tac ctxt Codegen.evaluation_conv;
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff changeset
   326
end