src/ZF/ex/LList.ML
changeset 11316 b4e71bd751e4
parent 11233 34c81a796ee3
equal deleted inserted replaced
11315:fbca0f74bcef 11316:b4e71bd751e4
    14 	Union_least, UN_least, Un_least, 
    14 	Union_least, UN_least, Un_least, 
    15 	Inter_greatest, Int_greatest, RepFun_subset,
    15 	Inter_greatest, Int_greatest, RepFun_subset,
    16 	Un_upper1, Un_upper2, Int_lower1, Int_lower2];
    16 	Un_upper1, Un_upper2, Int_lower1, Int_lower2];
    17 
    17 
    18 (*An elimination rule, for type-checking*)
    18 (*An elimination rule, for type-checking*)
    19 val LConsE = llist.mk_cases "LCons(a,l) : llist(A)";
    19 val LConsE = llist.mk_cases "LCons(a,l) \\<in> llist(A)";
    20 
    20 
    21 (*Proving freeness results*)
    21 (*Proving freeness results*)
    22 val LCons_iff      = llist.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'";
    22 val LCons_iff      = llist.mk_free "LCons(a,l)=LCons(a',l') <-> a=a' & l=l'";
    23 val LNil_LCons_iff = llist.mk_free "~ LNil=LCons(a,l)";
    23 val LNil_LCons_iff = llist.mk_free "~ LNil=LCons(a,l)";
    24 
    24 
    28 end;
    28 end;
    29 qed "llist_unfold";
    29 qed "llist_unfold";
    30 
    30 
    31 (*** Lemmas to justify using "llist" in other recursive type definitions ***)
    31 (*** Lemmas to justify using "llist" in other recursive type definitions ***)
    32 
    32 
    33 Goalw llist.defs "A<=B ==> llist(A) <= llist(B)";
    33 Goalw llist.defs "A \\<subseteq> B ==> llist(A) \\<subseteq> llist(B)";
    34 by (rtac gfp_mono 1);
    34 by (rtac gfp_mono 1);
    35 by (REPEAT (rtac llist.bnd_mono 1));
    35 by (REPEAT (rtac llist.bnd_mono 1));
    36 by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
    36 by (REPEAT (ares_tac (quniv_mono::basic_monos) 1));
    37 qed "llist_mono";
    37 qed "llist_mono";
    38 
    38 
    42                                  QPair_subset_univ,
    42                                  QPair_subset_univ,
    43                                  empty_subsetI, one_in_quniv RS qunivD];
    43                                  empty_subsetI, one_in_quniv RS qunivD];
    44 AddSDs [qunivD];
    44 AddSDs [qunivD];
    45 AddSEs [Ord_in_Ord];
    45 AddSEs [Ord_in_Ord];
    46 
    46 
    47 Goal "Ord(i) ==> ALL l: llist(quniv(A)). l Int Vset(i) <= univ(eclose(A))";
    47 Goal "Ord(i) ==> \\<forall>l \\<in> llist(quniv(A)). l Int Vset(i) \\<subseteq> univ(eclose(A))";
    48 by (etac trans_induct 1);
    48 by (etac trans_induct 1);
    49 by (rtac ballI 1);
    49 by (rtac ballI 1);
    50 by (etac llist.elim 1);
    50 by (etac llist.elim 1);
    51 by (rewrite_goals_tac ([QInl_def,QInr_def]@llist.con_defs));
    51 by (rewrite_goals_tac ([QInl_def,QInr_def]@llist.con_defs));
    52 (*LNil case*)
    52 (*LNil case*)
    53 by (Asm_simp_tac 1);
    53 by (Asm_simp_tac 1);
    54 (*LCons case*)
    54 (*LCons case*)
    55 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
    55 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
    56 qed "llist_quniv_lemma";
    56 qed "llist_quniv_lemma";
    57 
    57 
    58 Goal "llist(quniv(A)) <= quniv(A)";
    58 Goal "llist(quniv(A)) \\<subseteq> quniv(A)";
    59 by (rtac (qunivI RS subsetI) 1);
    59 by (rtac (qunivI RS subsetI) 1);
    60 by (rtac Int_Vset_subset 1);
    60 by (rtac Int_Vset_subset 1);
    61 by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1));
    61 by (REPEAT (ares_tac [llist_quniv_lemma RS bspec] 1));
    62 qed "llist_quniv";
    62 qed "llist_quniv";
    63 
    63 
    69 
    69 
    70 AddSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono];
    70 AddSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono];
    71 AddSEs [Ord_in_Ord, Pair_inject];
    71 AddSEs [Ord_in_Ord, Pair_inject];
    72 
    72 
    73 (*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
    73 (*Lemma for proving finality.  Unfold the lazy list; use induction hypothesis*)
    74 Goal "Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";
    74 Goal "Ord(i) ==> \\<forall>l l'. <l,l'> \\<in> lleq(A) --> l Int Vset(i) \\<subseteq> l'";
    75 by (etac trans_induct 1);
    75 by (etac trans_induct 1);
    76 by (REPEAT (resolve_tac [allI, impI] 1));
    76 by (REPEAT (resolve_tac [allI, impI] 1));
    77 by (etac lleq.elim 1);
    77 by (etac lleq.elim 1);
    78 by (rewrite_goals_tac (QInr_def::llist.con_defs));
    78 by (rewrite_goals_tac (QInr_def::llist.con_defs));
    79 by Safe_tac;
    79 by Safe_tac;
    83 bind_thm ("lleq_Int_Vset_subset",
    83 bind_thm ("lleq_Int_Vset_subset",
    84         (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp));
    84         (lleq_Int_Vset_subset_lemma RS spec RS spec RS mp));
    85 
    85 
    86 
    86 
    87 (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
    87 (*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
    88 val [prem] = goal LList.thy "<l,l'> : lleq(A) ==> <l',l> : lleq(A)";
    88 val [prem] = goal LList.thy "<l,l'> \\<in> lleq(A) ==> <l',l> \\<in> lleq(A)";
    89 by (rtac (prem RS converseI RS lleq.coinduct) 1);
    89 by (rtac (prem RS converseI RS lleq.coinduct) 1);
    90 by (rtac (lleq.dom_subset RS converse_type) 1);
    90 by (rtac (lleq.dom_subset RS converse_type) 1);
    91 by Safe_tac;
    91 by Safe_tac;
    92 by (etac lleq.elim 1);
    92 by (etac lleq.elim 1);
    93 by (ALLGOALS Fast_tac);
    93 by (ALLGOALS Fast_tac);
    94 qed "lleq_symmetric";
    94 qed "lleq_symmetric";
    95 
    95 
    96 Goal "<l,l'> : lleq(A) ==> l=l'";
    96 Goal "<l,l'> \\<in> lleq(A) ==> l=l'";
    97 by (rtac equalityI 1);
    97 by (rtac equalityI 1);
    98 by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
    98 by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1
    99      ORELSE etac lleq_symmetric 1));
    99      ORELSE etac lleq_symmetric 1));
   100 qed "lleq_implies_equal";
   100 qed "lleq_implies_equal";
   101 
   101 
   102 val [eqprem,lprem] = goal LList.thy
   102 val [eqprem,lprem] = goal LList.thy
   103     "[| l=l';  l: llist(A) |] ==> <l,l'> : lleq(A)";
   103     "[| l=l';  l \\<in> llist(A) |] ==> <l,l'> \\<in> lleq(A)";
   104 by (res_inst_tac [("X", "{<l,l>. l: llist(A)}")] lleq.coinduct 1);
   104 by (res_inst_tac [("X", "{<l,l>. l \\<in> llist(A)}")] lleq.coinduct 1);
   105 by (rtac (lprem RS RepFunI RS (eqprem RS subst)) 1);
   105 by (rtac (lprem RS RepFunI RS (eqprem RS subst)) 1);
   106 by Safe_tac;
   106 by Safe_tac;
   107 by (etac llist.elim 1);
   107 by (etac llist.elim 1);
   108 by (ALLGOALS Fast_tac);
   108 by (ALLGOALS Fast_tac);
   109 qed "equal_llist_implies_leq";
   109 qed "equal_llist_implies_leq";
   128 
   128 
   129 val lconst_subset = lconst_def RS def_lfp_subset;
   129 val lconst_subset = lconst_def RS def_lfp_subset;
   130 
   130 
   131 bind_thm ("member_subset_Union_eclose", (arg_into_eclose RS Union_upper));
   131 bind_thm ("member_subset_Union_eclose", (arg_into_eclose RS Union_upper));
   132 
   132 
   133 Goal "a : A ==> lconst(a) : quniv(A)";
   133 Goal "a \\<in> A ==> lconst(a) \\<in> quniv(A)";
   134 by (rtac (lconst_subset RS subset_trans RS qunivI) 1);
   134 by (rtac (lconst_subset RS subset_trans RS qunivI) 1);
   135 by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1);
   135 by (etac (arg_into_eclose RS eclose_subset RS univ_mono) 1);
   136 qed "lconst_in_quniv";
   136 qed "lconst_in_quniv";
   137 
   137 
   138 Goal "a:A ==> lconst(a): llist(A)";
   138 Goal "a \\<in> A ==> lconst(a): llist(A)";
   139 by (rtac (singletonI RS llist.coinduct) 1);
   139 by (rtac (singletonI RS llist.coinduct) 1);
   140 by (etac (lconst_in_quniv RS singleton_subsetI) 1);
   140 by (etac (lconst_in_quniv RS singleton_subsetI) 1);
   141 by (fast_tac (claset() addSIs [lconst]) 1);
   141 by (fast_tac (claset() addSIs [lconst]) 1);
   142 qed "lconst_type";
   142 qed "lconst_type";
   143 
   143 
   144 (*** flip --- equations merely assumed; certain consequences proved ***)
   144 (*** flip --- equations merely assumed; certain consequences proved ***)
   145 
   145 
   146 Addsimps [flip_LNil, flip_LCons, not_type];
   146 Addsimps [flip_LNil, flip_LCons, not_type];
   147 
   147 
   148 goal QUniv.thy "!!b. b:bool ==> b Int X <= univ(eclose(A))";
   148 goal QUniv.thy "!!b. b \\<in> bool ==> b Int X \\<subseteq> univ(eclose(A))";
   149 by (fast_tac (claset() addIs [Int_lower1 RS subset_trans] addSEs [boolE]) 1);
   149 by (fast_tac (claset() addIs [Int_lower1 RS subset_trans] addSEs [boolE]) 1);
   150 qed "bool_Int_subset_univ";
   150 qed "bool_Int_subset_univ";
   151 
   151 
   152 AddSIs [not_type];
   152 AddSIs [not_type];
   153 AddIs  [bool_Int_subset_univ];
   153 AddIs  [bool_Int_subset_univ];
   154 
   154 
   155 (*Reasoning borrowed from lleq.ML; a similar proof works for all
   155 (*Reasoning borrowed from lleq.ML; a similar proof works for all
   156   "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
   156   "productive" functions -- cf Coquand's "Infinite Objects in Type Theory".*)
   157 Goal "Ord(i) ==> ALL l: llist(bool). flip(l) Int Vset(i) <= \
   157 Goal "Ord(i) ==> \\<forall>l \\<in> llist(bool). flip(l) Int Vset(i) \\<subseteq> \
   158 \                   univ(eclose(bool))";
   158 \                   univ(eclose(bool))";
   159 by (etac trans_induct 1);
   159 by (etac trans_induct 1);
   160 by (rtac ballI 1);
   160 by (rtac ballI 1);
   161 by (etac llist.elim 1);
   161 by (etac llist.elim 1);
   162 by (ALLGOALS Asm_simp_tac);
   162 by (ALLGOALS Asm_simp_tac);
   164     (asm_simp_tac (simpset() addsimps [QInl_def,QInr_def] @ llist.con_defs)));
   164     (asm_simp_tac (simpset() addsimps [QInl_def,QInr_def] @ llist.con_defs)));
   165 (*LCons case*)
   165 (*LCons case*)
   166 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
   166 by (deepen_tac (claset() addIs [Ord_trans, Int_lower1 RS subset_trans]) 2 1);
   167 qed "flip_llist_quniv_lemma";
   167 qed "flip_llist_quniv_lemma";
   168 
   168 
   169 Goal "l: llist(bool) ==> flip(l) : quniv(bool)";
   169 Goal "l \\<in> llist(bool) ==> flip(l) \\<in> quniv(bool)";
   170 by (rtac (flip_llist_quniv_lemma RS bspec RS Int_Vset_subset RS qunivI) 1);
   170 by (rtac (flip_llist_quniv_lemma RS bspec RS Int_Vset_subset RS qunivI) 1);
   171 by (REPEAT (assume_tac 1));
   171 by (REPEAT (assume_tac 1));
   172 qed "flip_in_quniv";
   172 qed "flip_in_quniv";
   173 
   173 
   174 val [prem] = goal LList.thy "l : llist(bool) ==> flip(l): llist(bool)";
   174 val [prem] = goal LList.thy "l \\<in> llist(bool) ==> flip(l): llist(bool)";
   175 by (res_inst_tac [("X", "{flip(l) . l:llist(bool)}")]
   175 by (res_inst_tac [("X", "{flip(l) . l \\<in> llist(bool)}")]
   176        llist.coinduct 1);
   176        llist.coinduct 1);
   177 by (rtac (prem RS RepFunI) 1);
   177 by (rtac (prem RS RepFunI) 1);
   178 by (fast_tac (claset() addSIs [flip_in_quniv]) 1);
   178 by (fast_tac (claset() addSIs [flip_in_quniv]) 1);
   179 by (etac RepFunE 1);
   179 by (etac RepFunE 1);
   180 by (etac llist.elim 1);
   180 by (etac llist.elim 1);
   181 by (ALLGOALS Asm_simp_tac);
   181 by (ALLGOALS Asm_simp_tac);
   182 by (Fast_tac 1);
   182 by (Fast_tac 1);
   183 qed "flip_type";
   183 qed "flip_type";
   184 
   184 
   185 val [prem] = goal LList.thy
   185 val [prem] = goal LList.thy
   186     "l : llist(bool) ==> flip(flip(l)) = l";
   186     "l \\<in> llist(bool) ==> flip(flip(l)) = l";
   187 by (res_inst_tac [("X1", "{<flip(flip(l)),l> . l:llist(bool)}")]
   187 by (res_inst_tac [("X1", "{<flip(flip(l)),l> . l \\<in> llist(bool)}")]
   188        (lleq.coinduct RS lleq_implies_equal) 1);
   188        (lleq.coinduct RS lleq_implies_equal) 1);
   189 by (rtac (prem RS RepFunI) 1);
   189 by (rtac (prem RS RepFunI) 1);
   190 by (fast_tac (claset() addSIs [flip_type]) 1);
   190 by (fast_tac (claset() addSIs [flip_type]) 1);
   191 by (etac RepFunE 1);
   191 by (etac RepFunE 1);
   192 by (etac llist.elim 1);
   192 by (etac llist.elim 1);