QuotMain.thy
changeset 488 508f3106b89c
parent 487 f5db9ede89b0
child 489 2b7b349e470f
equal deleted inserted replaced
487:f5db9ede89b0 488:508f3106b89c
     2 imports QuotScript QuotList Prove
     2 imports QuotScript QuotList Prove
     3 uses ("quotient_info.ML") 
     3 uses ("quotient_info.ML") 
     4      ("quotient.ML")
     4      ("quotient.ML")
     5      ("quotient_def.ML")
     5      ("quotient_def.ML")
     6 begin
     6 begin
     7 
       
     8 ML {*
       
     9 let
       
    10    val parser = Args.context -- Scan.lift Args.name_source
       
    11    fun term_pat (ctxt, str) =
       
    12       str |> ProofContext.read_term_pattern ctxt
       
    13           |> ML_Syntax.print_term
       
    14           |> ML_Syntax.atomic
       
    15 in
       
    16    ML_Antiquote.inline "term_pat" (parser >> term_pat)
       
    17 end
       
    18 *}
       
    19 
       
    20 ML {*
       
    21 fun make_inst lhs t =
       
    22   let
       
    23     val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
       
    24     val _ $ (Abs (_, _, g)) = t;
       
    25     
       
    26   fun mk_abs i t =
       
    27       if incr_boundvars i u aconv t then Bound i
       
    28       else (case t of
       
    29         t1 $ t2 => mk_abs i t1 $ mk_abs i t2
       
    30       | Abs (s, T, t') => Abs (s, T, mk_abs (i + 1) t')
       
    31       | Bound j => if i = j then error "make_inst" else t
       
    32       | _ => t);
       
    33 
       
    34   in (f, Abs ("x", T, mk_abs 0 g)) end;
       
    35 *}
       
    36 
       
    37 ML {* 
       
    38   val trm1 = @{term_pat "(P::(nat\<Rightarrow>bool)\<Rightarrow>bool) (\<lambda>(x::nat). (?f ((g::nat\<Rightarrow>nat) x)))"}; 
       
    39   val trm2 = @{term_pat "(P::(nat\<Rightarrow>bool)\<Rightarrow>bool) (\<lambda>(x::nat). (h (g x) ((g::nat\<Rightarrow>nat) x)))"}
       
    40 *}
       
    41 
       
    42 ML {*
       
    43 fun test trm =
       
    44   case trm of
       
    45   (_ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u))) => (1,f)
       
    46   | (_ $ (Abs (_, _, (f as (Var (_, _)) $ u)))) => (2,f)
       
    47   | (_ $ (Abs (_, _, (f $ u)))) => (3,f)
       
    48 *}
       
    49 
       
    50 ML {* test trm1 *}
       
    51 
       
    52 ML {*
       
    53  make_inst trm1 trm2
       
    54   |> pairself (Syntax.string_of_term @{context})
       
    55   |> pairself writeln
       
    56 *}
       
    57 
       
    58 
       
    59 
       
    60 
     7 
    61 locale QUOT_TYPE =
     8 locale QUOT_TYPE =
    62   fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
     9   fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
    63   and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
    10   and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
    64   and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
    11   and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"