src/ZF/Constructible/Rec_Separation.thy
author wenzelm
Thu, 26 Apr 2007 15:41:49 +0200
changeset 22813 882513df2472
parent 21404 eb85850d3eb7
child 29223 e09c53289830
permissions -rw-r--r--
updated;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
     1
(*  Title:      ZF/Constructible/Rec_Separation.thy
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13566
diff changeset
     2
    ID:   $Id$
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
     4
*)
13429
wenzelm
parents: 13428
diff changeset
     5
wenzelm
parents: 13428
diff changeset
     6
header {*Separation for Facts About Recursion*}
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15766
diff changeset
     8
theory Rec_Separation imports Separation Internalize begin
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
     9
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    10
text{*This theory proves all instances needed for locales @{text
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13566
diff changeset
    11
"M_trancl"} and @{text "M_datatypes"}*}
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    12
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    13
lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> j<i"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    14
by simp
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    15
13493
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
    16
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    17
subsection{*The Locale @{text "M_trancl"}*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    18
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    19
subsubsection{*Separation for Reflexive/Transitive Closure*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    20
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    21
text{*First, The Defining Formula*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    22
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    23
(* "rtran_closure_mem(M,A,r,p) ==
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    24
      \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    25
       omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    26
       (\<exists>f[M]. typed_function(M,n',A,f) &
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    27
        (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    28
          fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    29
        (\<forall>j[M]. j\<in>n -->
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    30
          (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    31
            fun_apply(M,f,j,fj) & successor(M,j,sj) &
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    32
            fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    33
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    34
  rtran_closure_mem_fm :: "[i,i,i]=>i" where
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    35
 "rtran_closure_mem_fm(A,r,p) ==
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    36
   Exists(Exists(Exists(
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    37
    And(omega_fm(2),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    38
     And(Member(1,2),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    39
      And(succ_fm(1,0),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    40
       Exists(And(typed_function_fm(1, A#+4, 0),
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    41
        And(Exists(Exists(Exists(
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    42
              And(pair_fm(2,1,p#+7),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    43
               And(empty_fm(0),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    44
                And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    45
            Forall(Implies(Member(0,3),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    46
             Exists(Exists(Exists(Exists(
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    47
              And(fun_apply_fm(5,4,3),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    48
               And(succ_fm(4,2),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    49
                And(fun_apply_fm(5,2,1),
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    50
                 And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    51
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    52
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    53
lemma rtran_closure_mem_type [TC]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    54
 "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> rtran_closure_mem_fm(x,y,z) \<in> formula"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    55
by (simp add: rtran_closure_mem_fm_def)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    56
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    57
lemma sats_rtran_closure_mem_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    58
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    59
    ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <->
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    60
        rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    61
by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    62
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    63
lemma rtran_closure_mem_iff_sats:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    64
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    65
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    66
       ==> rtran_closure_mem(##A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    67
by (simp add: sats_rtran_closure_mem_fm)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    68
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
    69
lemma rtran_closure_mem_reflection:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    70
     "REFLECTS[\<lambda>x. rtran_closure_mem(L,f(x),g(x),h(x)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    71
               \<lambda>i x. rtran_closure_mem(##Lset(i),f(x),g(x),h(x))]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
    72
apply (simp only: rtran_closure_mem_def)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    73
apply (intro FOL_reflections function_reflections fun_plus_reflections)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    74
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    75
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    76
text{*Separation for @{term "rtrancl(r)"}.*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    77
lemma rtrancl_separation:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    78
     "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))"
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
    79
apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"],
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
    80
       auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
    81
apply (rule_tac env="[r,A]" in DPow_LsetI)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
    82
apply (rule rtran_closure_mem_iff_sats sep_rules | simp)+
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    83
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    84
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    85
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    86
subsubsection{*Reflexive/Transitive Closure, Internalized*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    87
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    88
(*  "rtran_closure(M,r,s) ==
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    89
        \<forall>A[M]. is_field(M,r,A) -->
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
    90
         (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    91
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    92
  rtran_closure_fm :: "[i,i]=>i" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    93
  "rtran_closure_fm(r,s) ==
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    94
   Forall(Implies(field_fm(succ(r),0),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    95
                  Forall(Iff(Member(0,succ(succ(s))),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    96
                             rtran_closure_mem_fm(1,succ(succ(r)),0)))))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    97
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    98
lemma rtran_closure_type [TC]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
    99
     "[| x \<in> nat; y \<in> nat |] ==> rtran_closure_fm(x,y) \<in> formula"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   100
by (simp add: rtran_closure_fm_def)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   101
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   102
lemma sats_rtran_closure_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   103
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   104
    ==> sats(A, rtran_closure_fm(x,y), env) <->
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   105
        rtran_closure(##A, nth(x,env), nth(y,env))"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   106
by (simp add: rtran_closure_fm_def rtran_closure_def)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   107
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   108
lemma rtran_closure_iff_sats:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   109
      "[| nth(i,env) = x; nth(j,env) = y;
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   110
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   111
       ==> rtran_closure(##A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   112
by simp
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   113
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   114
theorem rtran_closure_reflection:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   115
     "REFLECTS[\<lambda>x. rtran_closure(L,f(x),g(x)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   116
               \<lambda>i x. rtran_closure(##Lset(i),f(x),g(x))]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   117
apply (simp only: rtran_closure_def)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   118
apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   119
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   120
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   121
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   122
subsubsection{*Transitive Closure of a Relation, Internalized*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   123
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   124
(*  "tran_closure(M,r,t) ==
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   125
         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   126
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   127
  tran_closure_fm :: "[i,i]=>i" where
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   128
  "tran_closure_fm(r,s) ==
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   129
   Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   130
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   131
lemma tran_closure_type [TC]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   132
     "[| x \<in> nat; y \<in> nat |] ==> tran_closure_fm(x,y) \<in> formula"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   133
by (simp add: tran_closure_fm_def)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   134
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   135
lemma sats_tran_closure_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   136
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   137
    ==> sats(A, tran_closure_fm(x,y), env) <->
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   138
        tran_closure(##A, nth(x,env), nth(y,env))"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   139
by (simp add: tran_closure_fm_def tran_closure_def)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   140
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   141
lemma tran_closure_iff_sats:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   142
      "[| nth(i,env) = x; nth(j,env) = y;
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   143
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   144
       ==> tran_closure(##A, x, y) <-> sats(A, tran_closure_fm(i,j), env)"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   145
by simp
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   146
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   147
theorem tran_closure_reflection:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   148
     "REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   149
               \<lambda>i x. tran_closure(##Lset(i),f(x),g(x))]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   150
apply (simp only: tran_closure_def)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   151
apply (intro FOL_reflections function_reflections
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   152
             rtran_closure_reflection composition_reflection)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   153
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   154
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   155
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   156
subsubsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*}
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   157
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   158
lemma wellfounded_trancl_reflects:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   159
  "REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   160
                 w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp,
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   161
   \<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   162
       w \<in> Z & pair(##Lset(i),w,x,wx) & tran_closure(##Lset(i),r,rp) &
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   163
       wx \<in> rp]"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   164
by (intro FOL_reflections function_reflections fun_plus_reflections
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   165
          tran_closure_reflection)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   166
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   167
lemma wellfounded_trancl_separation:
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   168
         "[| L(r); L(Z) |] ==>
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   169
          separation (L, \<lambda>x.
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   170
              \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   171
               w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)"
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   172
apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"],
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   173
       auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   174
apply (rule_tac env="[r,Z]" in DPow_LsetI)
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   175
apply (rule sep_rules tran_closure_iff_sats | simp)+
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   176
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   177
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   178
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   179
subsubsection{*Instantiating the locale @{text M_trancl}*}
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   180
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   181
lemma M_trancl_axioms_L: "M_trancl_axioms(L)"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   182
  apply (rule M_trancl_axioms.intro)
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   183
   apply (assumption | rule rtrancl_separation wellfounded_trancl_separation)+
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   184
  done
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   185
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   186
theorem M_trancl_L: "PROP M_trancl(L)"
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   187
by (rule M_trancl.intro [OF M_basic_L M_trancl_axioms_L])
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   188
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   189
interpretation M_trancl [L] by (rule M_trancl_L)
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   190
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   191
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   192
subsection{*@{term L} is Closed Under the Operator @{term list}*}
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   193
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   194
subsubsection{*Instances of Replacement for Lists*}
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   195
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   196
lemma list_replacement1_Reflects:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   197
 "REFLECTS
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   198
   [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   199
         is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   200
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   201
         is_wfrec(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   202
                  iterates_MH(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   203
                          is_list_functor(##Lset(i), A), 0), memsn, u, y))]"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   204
by (intro FOL_reflections function_reflections is_wfrec_reflection
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   205
          iterates_MH_reflection list_functor_reflection)
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   206
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   207
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   208
lemma list_replacement1:
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   209
   "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   210
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   211
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   212
apply (rule_tac u="{B,A,n,0,Memrel(succ(n))}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   213
         in gen_separation_multi [OF list_replacement1_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   214
       auto simp add: nonempty)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   215
apply (rule_tac env="[B,A,n,0,Memrel(succ(n))]" in DPow_LsetI)
13434
78b93a667c01 better sats rules for higher-order operators
paulson
parents: 13429
diff changeset
   216
apply (rule sep_rules is_nat_case_iff_sats list_functor_iff_sats
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   217
            is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   218
done
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   219
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   220
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   221
lemma list_replacement2_Reflects:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   222
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   223
   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   224
                is_iterates(L, is_list_functor(L, A), 0, u, x),
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   225
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   226
               is_iterates(##Lset(i), is_list_functor(##Lset(i), A), 0, u, x)]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   227
by (intro FOL_reflections 
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   228
          is_iterates_reflection list_functor_reflection)
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   229
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   230
lemma list_replacement2:
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   231
   "L(A) ==> strong_replacement(L,
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   232
         \<lambda>n y. n\<in>nat & is_iterates(L, is_list_functor(L,A), 0, n, y))"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   233
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   234
apply (rule_tac u="{A,B,0,nat}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   235
         in gen_separation_multi [OF list_replacement2_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   236
       auto simp add: L_nat nonempty)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   237
apply (rule_tac env="[A,B,0,nat]" in DPow_LsetI)
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   238
apply (rule sep_rules list_functor_iff_sats is_iterates_iff_sats | simp)+
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   239
done
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   240
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   241
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   242
subsection{*@{term L} is Closed Under the Operator @{term formula}*}
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   243
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   244
subsubsection{*Instances of Replacement for Formulas*}
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   245
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   246
(*FIXME: could prove a lemma iterates_replacementI to eliminate the 
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   247
need to expand iterates_replacement and wfrec_replacement*)
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   248
lemma formula_replacement1_Reflects:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   249
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   250
   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   251
         is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   252
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   253
         is_wfrec(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   254
                  iterates_MH(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   255
                          is_formula_functor(##Lset(i)), 0), memsn, u, y))]"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   256
by (intro FOL_reflections function_reflections is_wfrec_reflection
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   257
          iterates_MH_reflection formula_functor_reflection)
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   258
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   259
lemma formula_replacement1:
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   260
   "iterates_replacement(L, is_formula_functor(L), 0)"
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   261
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   262
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   263
apply (rule_tac u="{B,n,0,Memrel(succ(n))}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   264
         in gen_separation_multi [OF formula_replacement1_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   265
       auto simp add: nonempty)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   266
apply (rule_tac env="[n,B,0,Memrel(succ(n))]" in DPow_LsetI)
13434
78b93a667c01 better sats rules for higher-order operators
paulson
parents: 13429
diff changeset
   267
apply (rule sep_rules is_nat_case_iff_sats formula_functor_iff_sats
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   268
            is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   269
done
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   270
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   271
lemma formula_replacement2_Reflects:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   272
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   273
   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   274
                is_iterates(L, is_formula_functor(L), 0, u, x),
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   275
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   276
               is_iterates(##Lset(i), is_formula_functor(##Lset(i)), 0, u, x)]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   277
by (intro FOL_reflections 
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   278
          is_iterates_reflection formula_functor_reflection)
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   279
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   280
lemma formula_replacement2:
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   281
   "strong_replacement(L,
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   282
         \<lambda>n y. n\<in>nat & is_iterates(L, is_formula_functor(L), 0, n, y))"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   283
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   284
apply (rule_tac u="{B,0,nat}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   285
         in gen_separation_multi [OF formula_replacement2_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   286
       auto simp add: nonempty L_nat)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   287
apply (rule_tac env="[B,0,nat]" in DPow_LsetI)
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   288
apply (rule sep_rules formula_functor_iff_sats is_iterates_iff_sats | simp)+
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   289
done
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   290
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   291
text{*NB The proofs for type @{term formula} are virtually identical to those
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   292
for @{term "list(A)"}.  It was a cut-and-paste job! *}
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   293
13387
b7464ca2ebbb new theorems to support Constructible proofs
paulson
parents: 13386
diff changeset
   294
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   295
subsubsection{*The Formula @{term is_nth}, Internalized*}
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   296
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   297
(* "is_nth(M,n,l,Z) ==
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   298
      \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   299
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   300
  nth_fm :: "[i,i,i]=>i" where
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   301
    "nth_fm(n,l,Z) == 
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   302
       Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0), 
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   303
              hd_fm(0,succ(Z))))"
13493
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   304
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   305
lemma nth_fm_type [TC]:
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   306
 "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> nth_fm(x,y,z) \<in> formula"
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   307
by (simp add: nth_fm_def)
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   308
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   309
lemma sats_nth_fm [simp]:
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   310
   "[| x < length(env); y \<in> nat; z \<in> nat; env \<in> list(A)|]
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   311
    ==> sats(A, nth_fm(x,y,z), env) <->
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   312
        is_nth(##A, nth(x,env), nth(y,env), nth(z,env))"
13493
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   313
apply (frule lt_length_in_nat, assumption)  
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   314
apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm) 
13493
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   315
done
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   316
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   317
lemma nth_iff_sats:
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   318
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z;
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   319
          i < length(env); j \<in> nat; k \<in> nat; env \<in> list(A)|]
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   320
       ==> is_nth(##A, x, y, z) <-> sats(A, nth_fm(i,j,k), env)"
13493
5aa68c051725 Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents: 13441
diff changeset
   321
by (simp add: sats_nth_fm)
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   322
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   323
theorem nth_reflection:
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   324
     "REFLECTS[\<lambda>x. is_nth(L, f(x), g(x), h(x)),  
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   325
               \<lambda>i x. is_nth(##Lset(i), f(x), g(x), h(x))]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   326
apply (simp only: is_nth_def)
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   327
apply (intro FOL_reflections is_iterates_reflection
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   328
             hd_reflection tl_reflection) 
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   329
done
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   330
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   331
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   332
subsubsection{*An Instance of Replacement for @{term nth}*}
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   333
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   334
(*FIXME: could prove a lemma iterates_replacementI to eliminate the 
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   335
need to expand iterates_replacement and wfrec_replacement*)
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   336
lemma nth_replacement_Reflects:
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   337
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   338
   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   339
         is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   340
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   341
         is_wfrec(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   342
                  iterates_MH(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   343
                          is_tl(##Lset(i)), z), memsn, u, y))]"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   344
by (intro FOL_reflections function_reflections is_wfrec_reflection
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   345
          iterates_MH_reflection tl_reflection)
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   346
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   347
lemma nth_replacement:
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   348
   "L(w) ==> iterates_replacement(L, is_tl(L), w)"
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   349
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   350
apply (rule strong_replacementI)
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   351
apply (rule_tac u="{B,w,Memrel(succ(n))}" 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   352
         in gen_separation_multi [OF nth_replacement_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   353
       auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   354
apply (rule_tac env="[B,w,Memrel(succ(n))]" in DPow_LsetI)
13434
78b93a667c01 better sats rules for higher-order operators
paulson
parents: 13429
diff changeset
   355
apply (rule sep_rules is_nat_case_iff_sats tl_iff_sats
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   356
            is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   357
done
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   358
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   359
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   360
subsubsection{*Instantiating the locale @{text M_datatypes}*}
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   361
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   362
lemma M_datatypes_axioms_L: "M_datatypes_axioms(L)"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   363
  apply (rule M_datatypes_axioms.intro)
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   364
      apply (assumption | rule
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   365
        list_replacement1 list_replacement2
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   366
        formula_replacement1 formula_replacement2
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   367
        nth_replacement)+
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   368
  done
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   369
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   370
theorem M_datatypes_L: "PROP M_datatypes(L)"
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   371
  apply (rule M_datatypes.intro)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   372
   apply (rule M_trancl_L)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   373
  apply (rule M_datatypes_axioms_L) 
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   374
  done
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   375
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   376
interpretation M_datatypes [L] by (rule M_datatypes_L)
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   377
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   378
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   379
subsection{*@{term L} is Closed Under the Operator @{term eclose}*}
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   380
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   381
subsubsection{*Instances of Replacement for @{term eclose}*}
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   382
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   383
lemma eclose_replacement1_Reflects:
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   384
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   385
   [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) &
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   386
         is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)),
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   387
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   388
         is_wfrec(##Lset(i),
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   389
                  iterates_MH(##Lset(i), big_union(##Lset(i)), A),
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   390
                  memsn, u, y))]"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   391
by (intro FOL_reflections function_reflections is_wfrec_reflection
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   392
          iterates_MH_reflection)
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   393
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   394
lemma eclose_replacement1:
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   395
   "L(A) ==> iterates_replacement(L, big_union(L), A)"
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   396
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   397
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   398
apply (rule_tac u="{B,A,n,Memrel(succ(n))}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   399
         in gen_separation_multi [OF eclose_replacement1_Reflects], auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   400
apply (rule_tac env="[B,A,n,Memrel(succ(n))]" in DPow_LsetI)
13434
78b93a667c01 better sats rules for higher-order operators
paulson
parents: 13429
diff changeset
   401
apply (rule sep_rules iterates_MH_iff_sats is_nat_case_iff_sats
13441
d6d620639243 better satisfies rules for is_recfun
paulson
parents: 13440
diff changeset
   402
             is_wfrec_iff_sats big_union_iff_sats quasinat_iff_sats | simp)+
13409
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   403
done
d4ea094c650e Relativization and Separation for the function "nth"
paulson
parents: 13398
diff changeset
   404
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   405
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   406
lemma eclose_replacement2_Reflects:
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   407
 "REFLECTS
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   408
   [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat &
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   409
                is_iterates(L, big_union(L), A, u, x),
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   410
    \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat &
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   411
               is_iterates(##Lset(i), big_union(##Lset(i)), A, u, x)]"
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   412
by (intro FOL_reflections function_reflections is_iterates_reflection)
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   413
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   414
lemma eclose_replacement2:
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   415
   "L(A) ==> strong_replacement(L,
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   416
         \<lambda>n y. n\<in>nat & is_iterates(L, big_union(L), A, n, y))"
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   417
apply (rule strong_replacementI)
13566
52a419210d5c Streamlined proofs of instances of Separation
paulson
parents: 13564
diff changeset
   418
apply (rule_tac u="{A,B,nat}" 
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   419
         in gen_separation_multi [OF eclose_replacement2_Reflects],
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   420
       auto simp add: L_nat)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13655
diff changeset
   421
apply (rule_tac env="[A,B,nat]" in DPow_LsetI)
13655
95b95cdb4704 Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents: 13651
diff changeset
   422
apply (rule sep_rules is_iterates_iff_sats big_union_iff_sats | simp)+
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   423
done
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   424
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   425
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   426
subsubsection{*Instantiating the locale @{text M_eclose}*}
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   427
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   428
lemma M_eclose_axioms_L: "M_eclose_axioms(L)"
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   429
  apply (rule M_eclose_axioms.intro)
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   430
   apply (assumption | rule eclose_replacement1 eclose_replacement2)+
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   431
  done
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   432
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   433
theorem M_eclose_L: "PROP M_eclose(L)"
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   434
  apply (rule M_eclose.intro)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   435
   apply (rule M_datatypes_L)
13437
01b3fc0cc1b8 separate "axioms" proofs: more flexible for locale reasoning
paulson
parents: 13434
diff changeset
   436
  apply (rule M_eclose_axioms_L)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13422
diff changeset
   437
  done
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   438
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   439
interpretation M_eclose [L] by (rule M_eclose_L)
15766
b08feb003f3c Cleaned up, now use interpretation.
ballarin
parents: 13807
diff changeset
   440
13422
af9bc8d87a75 Added the assumption nth_replacement to locale M_datatypes.
paulson
parents: 13418
diff changeset
   441
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents:
diff changeset
   442
end