From 2b12e53cf4c4ff08773459ddbbb5d50749def0a4 Mon Sep 17 00:00:00 2001 From: Burkhart Wolff Date: Wed, 3 Apr 2024 20:51:21 +0200 Subject: [PATCH] first monitor proof without sorries. --- .../technical_report/technical_report.thy | 56 ++++++++----------- 1 file changed, 22 insertions(+), 34 deletions(-) diff --git a/Isabelle_DOF/ontologies/technical_report/technical_report.thy b/Isabelle_DOF/ontologies/technical_report/technical_report.thy index c6a64a8..19aebbf 100644 --- a/Isabelle_DOF/ontologies/technical_report/technical_report.thy +++ b/Isabelle_DOF/ontologies/technical_report/technical_report.thy @@ -81,8 +81,7 @@ doc_class report = bibliography ~~ \index\ ~~ \annex\\<^sup>* )" -ML\ML_Syntax.print_string\ -ML\ML_Syntax.atomic\ + (* ML\@{doc_class "report"}\ *) @@ -94,6 +93,8 @@ ML\ fun get_class_2_ML ctxt (str,_) = setup\ML_Antiquotation.inline \<^binding>\doc_class2\ (fn (ctxt,toks) => (AttributeAccess.parse_cid >> get_class_2_ML ctxt) (ctxt, toks))\ +ML\@{term \a + b\}\ + ML\@{doc_class2 "report"}\ ML\fun constant_def (decl, spec, prems, params) = #2 o Specification.definition decl params prems spec\ @@ -138,47 +139,34 @@ notation Plus (infixr "||" 55) notation Times (infixr "~~" 60) notation Atom ("\_\" 65) - -lemma seq_cancel : "Lang (a ~~ b) \ Lang (a ~~ c) = (Lang (b) \ Lang (c))" - apply auto - thm subsetCE - apply(drule_tac c=x in subsetCE) - prefer 3 apply assumption - find_theorems subset_eq elim - prefer 2 - sledgehammer - - sorry - - -lemma opt_star_incl:"Lang (opt a) \ Lang (Star a)" by (simp add: opt_def subset_iff) -lemma rep1_star_incl:"Lang (rep1 a) \ Lang (Star a)" - unfolding rep1_def - apply(subst L_Star) - apply(subst L_Conc) - by force - - -lemma regexp_unit_right : "Lang (a) = Lang (a ~~ One) " by simp -lemma regexp_unit_left : "Lang (a) = Lang (One ~~ a) " by simp - lemma regexp_seq_mono : "Lang(a) \ Lang (a') \ Lang(b) \ Lang (b') \ Lang(a ~~ b) \ Lang(a' ~~ b')" by auto -lemma seq_d1 : "Lang (opt b ~~ a) \ Lang ( c) = (Lang (a) \ Lang (c))" sorry -lemma seq_d2 : "Lang (rep1 b ~~ a) \ Lang ( c) = (Lang (a) \ Lang (c))" sorry -lemma seq_d3 : "Lang (a) \ Lang (opt b ~~ c) = (Lang (a) \ Lang (c))" sorry -lemma seq_d4 : "Lang (a) \ Lang (Star b ~~ c) = (Lang (a) \ Lang (c))" sorry +lemma regexp_unit_right : "Lang (a) = Lang (a ~~ One) " by simp +lemma regexp_unit_left : "Lang (a) = Lang (One ~~ a) " by simp + +lemma opt_star_incl :"Lang (opt a) \ Lang (Star a)" by (simp add: opt_def subset_iff) +lemma rep1_star_incl:"Lang (rep1 a) \ Lang (Star a)" + unfolding rep1_def by(subst L_Star, subst L_Conc)(force) +lemma cancel_rep1 : "Lang (a) \ Lang (rep1 a)" + unfolding rep1_def by auto + + +lemma seq_cancel_opt : "Lang (a) \ Lang (c) \ Lang (a) \ Lang (opt b ~~ c)" + by(subst regexp_unit_left, rule regexp_seq_mono)(simp_all add: opt_def) + +lemma seq_cancel_Star : "Lang (a) \ Lang (c) \ Lang (a) \ Lang (Star b ~~ c)" + by auto text\Not a terribly deep theorem, but an interesting property of consistency between ontologies - so, a lemma that shouldn't break if the involved ontologies were changed: -the structural language of articles should be included in the structural language of reports, -that is to say, reports should just have a richer structure than ordinary papers: \ +the structural language of articles should be included in the structural language of +reports, that is to say, reports should just have a richer structure than ordinary papers. \ -theorem articles_sub_reports: \(Lang article_monitor) \ Lang report_monitor\ +theorem articles_sub_reports: \(Lang article_monitor) \ Lang report_monitor\ unfolding article_monitor_def report_monitor_def - apply(simp only: seq_cancel seq_d1 seq_d2 seq_d3) + apply(rule regexp_seq_mono[OF subset_refl] | rule seq_cancel_opt | rule subset_refl)+ done