src/HOL/Library/reflection.ML
author Christian Sternagel
Thu, 13 Dec 2012 13:11:38 +0100
changeset 50516 ed6b40d15d1c
parent 46763 aa9f5c3bcd4c
child 51717 9e7d1c139569
permissions -rw-r--r--
renamed "emb" to "list_hembeq"; make "list_hembeq" reflexive independent of the base order; renamed "sub" to "sublisteq"; dropped "transp_on" (state transitivity explicitly instead); no need to hide "sub" after renaming; replaced some ASCII symbols by proper Isabelle symbols; NEWS
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29650
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 29273
diff changeset
     1
(*  Title:      HOL/Library/reflection.ML
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 29273
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     3
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     4
A trial for automatical reification.
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     5
*)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
     6
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 24630
diff changeset
     7
signature REFLECTION =
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 24630
diff changeset
     8
sig
20319
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
31810
a6b800855cdd Added new evaluator "approximate"
hoelzl
parents: 31794
diff changeset
    13
  val genreif : Proof.context -> thm list -> term -> thm
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    14
end;
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    15
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 24630
diff changeset
    16
structure Reflection : REFLECTION =
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 24630
diff changeset
    17
struct
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    18
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    19
  (* Make a congruence rule out of a defining equation for the interpretation *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    20
  (* th is one defining equation of f, i.e.
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    21
     th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    22
  (* Cp is a constructor pattern and P is a pattern *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    23
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    24
  (* The result is:
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    25
      [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) *)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    26
  (*  + 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
    27
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    28
fun mk_congeq ctxt fs th =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    29
  let
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
    30
   val Const (fN, _) = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
29650
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 29273
diff changeset
    31
     |> fst |> strip_comb |> fst
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42284
diff changeset
    32
   val thy = Proof_Context.theory_of ctxt
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    33
   val cert = Thm.cterm_of thy
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 31412
diff changeset
    34
   val (((_,_),[th']), ctxt') = Variable.import true [th] ctxt
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    35
   val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th'))
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    36
   fun add_fterms (t as t1 $ t2) =
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 24630
diff changeset
    37
       if exists (fn f => Term.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
    38
       else add_fterms t1 #> add_fterms t2
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
    39
     | add_fterms (t as Abs _) =
29273
285c00993bc2 use exists_Const directly;
wenzelm
parents: 29269
diff changeset
    40
       if exists_Const (fn (c, _) => c = fN) t then (fn _ => [t]) else (fn _ => [])
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    41
     | add_fterms _ = I
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    42
   val fterms = add_fterms rhs []
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff changeset
    43
   val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    44
   val tys = map fastype_of fterms
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    45
   val vs = map Free (xs ~~ tys)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    46
   val env = fterms ~~ vs
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    47
                    (* FIXME!!!!*)
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
    48
   fun replace_fterms (t as t1 $ t2) =
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    49
       (case AList.lookup (op aconv) env t of
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    50
            SOME v => v
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    51
          | NONE => replace_fterms t1 $ replace_fterms t2)
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    52
     | replace_fterms t = (case AList.lookup (op aconv) env t of
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    53
                               SOME v => v
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    54
                             | NONE => t)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    55
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    56
   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
    57
     | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t))
46510
696f3fec3f83 more antiquotations;
wenzelm
parents: 45403
diff changeset
    58
   fun tryext x = (x RS @{lemma "(\<forall>x. f x = g x) \<Longrightarrow> f = g" by blast} handle THM _ =>  x)
45403
wenzelm
parents: 43960
diff changeset
    59
   val cong =
wenzelm
parents: 43960
diff changeset
    60
    (Goal.prove ctxt'' [] (map mk_def env)
wenzelm
parents: 43960
diff changeset
    61
      (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs)))
wenzelm
parents: 43960
diff changeset
    62
      (fn {context, prems, ...} =>
wenzelm
parents: 43960
diff changeset
    63
        Local_Defs.unfold_tac context (map tryext prems) THEN rtac th' 1)) RS sym
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    64
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    65
   val (cong' :: vars') =
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    66
       Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o cert) vs)
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    67
   val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    68
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    69
  in  (vs', cong') end;
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    70
 (* 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
    71
        (* [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    72
val FWD = curry (op OF);
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    73
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
    74
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    75
exception REIF of string;
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
    76
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
    77
fun dest_listT (Type (@{type_name "list"}, [T])) = T;
29834
3237cfd177f3 Proof method 'reify' is now reentrant.
hoelzl
parents: 29805
diff changeset
    78
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    79
fun rearrange congs =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    80
  let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    81
    fun P (_, th) =
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38549
diff changeset
    82
      let val @{term "Trueprop"}$(Const (@{const_name HOL.eq},_) $l$_) = concl_of th
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    83
      in can dest_Var l end
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    84
    val (yes,no) = List.partition P congs
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    85
  in no @ yes end
29834
3237cfd177f3 Proof method 'reify' is now reentrant.
hoelzl
parents: 29805
diff changeset
    86
3237cfd177f3 Proof method 'reify' is now reentrant.
hoelzl
parents: 29805
diff changeset
    87
fun genreif ctxt raw_eqs t =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    88
  let
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
    89
    fun index_of t bds =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    90
      let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    91
        val tt = HOLogic.listT (fastype_of t)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    92
      in
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
    93
       (case AList.lookup Type.could_unify bds tt of
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    94
          NONE => error "index_of : type not found in environements!"
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    95
        | SOME (tbs,tats) =>
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
    96
          let
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31810
diff changeset
    97
            val i = find_index (fn t' => t' = t) tats
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31810
diff changeset
    98
            val j = find_index (fn t' => t' = t) tbs
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
    99
          in (if j = ~1 then
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   100
              if i = ~1
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   101
              then (length tbs + length tats,
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   102
                    AList.update Type.could_unify (tt,(tbs,tats@[t])) bds)
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   103
              else (i, bds) else (j, bds))
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   104
          end)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   105
      end;
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   106
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   107
    (* Generic decomp for reification : matches the actual term with the
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   108
       rhs of one cong rule. The result of the matching guides the
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   109
       proof synthesis: The matches of the introduced Variables A1 .. An are
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   110
       processed recursively
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   111
       The rest is instantiated in the cong rule,i.e. no reification is needed *)
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
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   113
    (* da is the decomposition for atoms, ie. it returns ([],g) where g
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   114
       returns the right instance f (AtC n) = t , where AtC is the Atoms
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   115
       constructor and n is the number of the atom corresponding to t *)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   116
    fun decomp_genreif da cgns (t,ctxt) bds =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   117
      let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42284
diff changeset
   118
        val thy = Proof_Context.theory_of ctxt
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   119
        val cert = cterm_of thy
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   120
        fun tryabsdecomp (s,ctxt) bds =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   121
          (case s of
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   122
             Abs(_, xT, ta) => (
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   123
               let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   124
                 val ([xn],ctxt') = Variable.variant_fixes ["x"] ctxt
42284
326f57825e1a explicit structure Syntax_Trans;
wenzelm
parents: 38864
diff changeset
   125
                 val (xn,ta) = Syntax_Trans.variant_abs (xn,xT,ta)  (* FIXME !? *)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   126
                 val x = Free(xn,xT)
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   127
                 val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   128
                          of NONE => error "tryabsdecomp: Type not found in the Environement"
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   129
                           | SOME (bsT,atsT) =>
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   130
                             (AList.update Type.could_unify (HOLogic.listT xT, ((x::bsT), atsT)) bds))
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   131
               in (([(ta, ctxt')],
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   132
                    fn ([th], bds) =>
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   133
                      (hd (Variable.export ctxt' ctxt [(Thm.forall_intr (cert x) th) COMP allI]),
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   134
                       let val (bsT,asT) = the(AList.lookup Type.could_unify bds (HOLogic.listT xT))
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   135
                       in AList.update Type.could_unify (HOLogic.listT xT,(tl bsT,asT)) bds
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   136
                       end)),
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   137
                   bds)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   138
               end)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   139
           | _ => da (s,ctxt) bds)
37117
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   140
      in
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   141
        (case cgns of
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   142
          [] => tryabsdecomp (t,ctxt) bds
37117
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   143
        | ((vns,cong)::congs) =>
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   144
            (let
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   145
              val cert = cterm_of thy
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   146
              val certy = ctyp_of thy
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   147
              val (tyenv, tmenv) =
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   148
                Pattern.match thy
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   149
                  ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t)
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   150
                  (Vartab.empty, Vartab.empty)
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   151
              val (fnvs,invs) = List.partition (fn ((vn,_),_) => member (op =) vns vn) (Vartab.dest tmenv)
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   152
              val (fts,its) =
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   153
                (map (snd o snd) fnvs,
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   154
                 map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t)) invs)
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   155
              val ctyenv = map (fn ((vn,vi),(s,ty)) => (certy (TVar((vn,vi),s)), certy ty)) (Vartab.dest tyenv)
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   156
            in ((fts ~~ (replicate (length fts) ctxt),
43333
2bdec7f430d3 renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents: 42426
diff changeset
   157
                 Library.apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds)
37117
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   158
            end handle Pattern.MATCH => decomp_genreif da congs (t,ctxt) bds))
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   159
      end;
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   160
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   161
 (* looks for the atoms equation and instantiates it with the right number *)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   162
    fun mk_decompatom eqs (t,ctxt) bds = (([], fn (_, bds) =>
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   163
      let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   164
        val tT = fastype_of t
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   165
        fun isat eq =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   166
          let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   167
            val rhs = eq |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   168
          in exists_Const
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   169
            (fn (n,ty) => n = @{const_name "List.nth"}
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   170
                          andalso
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   171
                          AList.defined Type.could_unify bds (domain_type ty)) rhs
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   172
            andalso Type.could_unify (fastype_of rhs, tT)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   173
          end
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   174
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   175
        fun get_nths t acc =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   176
          case t of
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   177
            Const(@{const_name "List.nth"},_)$vs$n => insert (fn ((a,_),(b,_)) => a aconv b) (t,(vs,n)) acc
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   178
          | t1$t2 => get_nths t1 (get_nths t2 acc)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   179
          | Abs(_,_,t') => get_nths t'  acc
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   180
          | _ => acc
23548
e25991f126ce Generalized case for atoms. Selection of environment lists is allowed more than once.
chaieb
parents: 22568
diff changeset
   181
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   182
        fun
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   183
           tryeqs [] bds = error "Can not find the atoms equation"
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   184
         | tryeqs (eq::eqs) bds = ((
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   185
          let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   186
            val rhs = eq |> prop_of |> HOLogic.dest_Trueprop  |> HOLogic.dest_eq |> snd
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   187
            val nths = get_nths rhs []
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   188
            val (vss,_ ) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) =>
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   189
              (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], [])
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   190
            val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   191
            val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt'
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42284
diff changeset
   192
            val thy = Proof_Context.theory_of ctxt''
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   193
            val cert = cterm_of thy
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   194
            val certT = ctyp_of thy
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   195
            val vsns_map = vss ~~ vsns
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   196
            val xns_map = (fst (split_list nths)) ~~ xns
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   197
            val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   198
            val rhs_P = subst_free subst rhs
32032
a6a6e8031c14 tuned/modernized Envir operations;
wenzelm
parents: 31986
diff changeset
   199
            val (tyenv, tmenv) = Pattern.match thy (rhs_P, t) (Vartab.empty, Vartab.empty)
32035
8e77b6a250d5 tuned/modernized Envir.subst_XXX;
wenzelm
parents: 32032
diff changeset
   200
            val sbst = Envir.subst_term (tyenv, tmenv)
8e77b6a250d5 tuned/modernized Envir.subst_XXX;
wenzelm
parents: 32032
diff changeset
   201
            val sbsT = Envir.subst_type tyenv
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   202
            val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   203
                               (Vartab.dest tyenv)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   204
            val tml = Vartab.dest tmenv
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   205
            val (subst_ns, bds) = fold_map
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   206
                (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds =>
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   207
                  let
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   208
                    val name = snd (the (AList.lookup (op =) tml xn0))
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   209
                    val (idx, bds) = index_of name bds
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   210
                  in ((cert n, idx |> (HOLogic.mk_nat #> cert)), bds) end) subst bds
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   211
            val subst_vs =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   212
              let
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   213
                fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   214
                  let
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   215
                    val cns = sbst (Const(@{const_name "List.Cons"}, T --> lT --> lT))
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   216
                    val lT' = sbsT lT
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   217
                    val (bsT, _) = the (AList.lookup Type.could_unify bds lT)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   218
                    val vsn = the (AList.lookup (op =) vsns_map vs)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   219
                    val cvs = cert (fold_rev (fn x => fn xs => cns$x$xs) bsT (Free (vsn, lT')))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   220
                  in (cert vs, cvs) end
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   221
              in map h subst end
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   222
            val cts = map (fn ((vn,vi),(tT,t)) => (cert(Var ((vn,vi),tT)), cert t))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   223
                          (fold (AList.delete (fn (((a: string),_),(b,_)) => a = b))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   224
                                (map (fn n => (n,0)) xns) tml)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   225
            val substt =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   226
              let val ih = Drule.cterm_rule (Thm.instantiate (subst_ty,[]))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   227
              in map (fn (v,t) => (ih v, ih t)) (subst_ns@subst_vs@cts)  end
43333
2bdec7f430d3 renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents: 42426
diff changeset
   228
            val th = (Drule.instantiate_normalize (subst_ty, substt)  eq) RS sym
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   229
          in (hd (Variable.export ctxt'' ctxt [th]), bds) end)
37117
59cee8807c29 eliminated various catch-all exception patterns, guessing at the concrete exeptions that are intended here;
wenzelm
parents: 36945
diff changeset
   230
          handle Pattern.MATCH => tryeqs eqs bds)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   231
      in tryeqs (filter isat eqs) bds end), bds);
20374
01b711328990 Reification now handels binders.
chaieb
parents: 20319
diff changeset
   232
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   233
  (* Generic reification procedure: *)
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   234
  (* 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
   235
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   236
    fun mk_congs ctxt raw_eqs =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   237
      let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   238
        val fs = fold_rev (fn eq =>
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   239
                           insert (op =) (eq |> prop_of |> HOLogic.dest_Trueprop
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   240
                           |> HOLogic.dest_eq |> fst |> strip_comb
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   241
                           |> fst)) raw_eqs []
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   242
        val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   243
                            ) fs []
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   244
        val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42284
diff changeset
   245
        val thy = Proof_Context.theory_of ctxt'
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   246
        val cert = cterm_of thy
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   247
        val vstys = map (fn (t,v) => (t,SOME (cert (Free(v,t)))))
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   248
                    (tys ~~ vs)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   249
        val is_Var = can dest_Var
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   250
        fun insteq eq vs =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   251
          let
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   252
            val subst = map (fn (v as Var(_, t)) => (cert v, (the o the) (AList.lookup (op =) vstys t)))
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   253
                        (filter is_Var vs)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   254
          in Thm.instantiate ([],subst) eq
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   255
          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
   256
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   257
        val bds = AList.make (fn _ => ([],[])) tys
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   258
        val eqs = map (fn eq => eq |> prop_of |> HOLogic.dest_Trueprop
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   259
                                   |> HOLogic.dest_eq |> fst |> strip_comb |> snd |> tl
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   260
                                   |> (insteq eq)) raw_eqs
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   261
        val (ps,congs) = split_list (map (mk_congeq ctxt' fs) eqs)
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   262
      in (ps ~~ (Variable.export ctxt' ctxt congs), bds)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   263
      end
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   264
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   265
    val (congs, bds) = mk_congs ctxt raw_eqs
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   266
    val congs = rearrange congs
31412
f2e6b6526092 Converted reification to use fold_map instead of Library.foldl_map. Use antiquotations.
hoelzl
parents: 31387
diff changeset
   267
    val (th, bds) = divide_and_conquer' (decomp_genreif (mk_decompatom raw_eqs) congs) (t,ctxt) bds
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   268
    fun is_listVar (Var (_,t)) = can dest_listT t
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   269
         | is_listVar _ = false
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   270
    val vars = th |> prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   271
                  |> strip_comb |> snd |> filter is_listVar
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42284
diff changeset
   272
    val cert = cterm_of (Proof_Context.theory_of ctxt)
46763
aa9f5c3bcd4c dropped dead code
haftmann
parents: 46510
diff changeset
   273
    val cvs = map (fn (v as Var(_, t)) => (cert v,
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 31386
diff changeset
   274
                  the (AList.lookup Type.could_unify bds t) |> snd |> HOLogic.mk_list (dest_listT t) |> cert)) vars
43333
2bdec7f430d3 renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents: 42426
diff changeset
   275
    val th' = Drule.instantiate_normalize ([], cvs) th
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   276
    val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th'
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   277
    val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t')))
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   278
               (fn _ => simp_tac (simpset_of ctxt) 1)
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   279
  in FWD trans [th'',th']
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   280
  end
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   281
23648
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   282
bccbf6138c30 Try several correctness theorems for reflection; rearrange cong rules to avoid the absoption cases;
chaieb
parents: 23643
diff changeset
   283
fun genreflect ctxt conv corr_thms raw_eqs t =
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   284
  let
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   285
    val reifth = genreif ctxt raw_eqs t
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   286
    fun trytrans [] = error "No suitable correctness theorem found"
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   287
      | trytrans (th::ths) =
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   288
           (FWD trans [reifth, th RS sym] handle THM _ => trytrans ths)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   289
    val th = trytrans corr_thms
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   290
    val ft = (Thm.dest_arg1 o Thm.dest_arg o Thm.dest_arg o cprop_of) th
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   291
    val rth = conv ft
46510
696f3fec3f83 more antiquotations;
wenzelm
parents: 45403
diff changeset
   292
  in
696f3fec3f83 more antiquotations;
wenzelm
parents: 45403
diff changeset
   293
    simplify (HOL_basic_ss addsimps raw_eqs addsimps @{thms nth_Cons_0 nth_Cons_Suc})
31386
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   294
             (simplify (HOL_basic_ss addsimps [rth]) th)
8624b75a7784 corrected spacing in reflection
hoelzl
parents: 30969
diff changeset
   295
  end
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   296
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   297
fun genreify_tac ctxt eqs to = SUBGOAL (fn (goal, i) =>
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   298
  let
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   299
    val t = (case to of NONE => HOLogic.dest_Trueprop goal | SOME x => x)
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   300
    val th = genreif ctxt eqs t RS ssubst
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   301
  in rtac th i end);
20319
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   302
a8761e8568de Generic reflection and reification (by Amine Chaieb).
wenzelm
parents:
diff changeset
   303
    (* Reflection calls reification and uses the correctness *)
43959
285ffb18da30 fixed typo
bulwahn
parents: 43333
diff changeset
   304
        (* theorem assumed to be the head of the list *)
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   305
fun gen_reflection_tac ctxt conv corr_thms raw_eqs to = SUBGOAL (fn (goal, i) =>
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   306
  let
42368
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   307
    val t = (case to of NONE => HOLogic.dest_Trueprop goal | SOME x => x)
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   308
    val th = genreflect ctxt conv corr_thms raw_eqs t RS ssubst
3b8498ac2314 proper subgoal addressing via SUBGOAL/CSUBGOAL -- assuming these tactics did not handle Subscript in any special way;
wenzelm
parents: 42361
diff changeset
   309
  in rtac th i THEN TRY (rtac TrueI i) end);  (* FIXME THEN_ALL_NEW !? *)
21878
cfc07819bb47 added gen_reflection_tac
haftmann
parents: 21757
diff changeset
   310
43960
c2554cc82d34 replacing conversion function of old code generator by the current code generator in the reflection tactic
bulwahn
parents: 43959
diff changeset
   311
fun reflection_tac ctxt = gen_reflection_tac ctxt
c2554cc82d34 replacing conversion function of old code generator by the current code generator in the reflection tactic
bulwahn
parents: 43959
diff changeset
   312
  (Code_Evaluation.dynamic_conv (Proof_Context.theory_of ctxt));
c2554cc82d34 replacing conversion function of old code generator by the current code generator in the reflection tactic
bulwahn
parents: 43959
diff changeset
   313
  (*FIXME why Code_Evaluation.dynamic_conv?  very specific...*)
29650
cc3958d31b1d Reflection.thy now in HOL/Library
haftmann
parents: 29273
diff changeset
   314
20797
c1f0bc7e7d80 renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents: 20595
diff changeset
   315
end